whnf actually reduces to whnf now (probably)
This commit is contained in:
parent
f097e1c091
commit
92617a2e4a
11 changed files with 693 additions and 679 deletions
|
@ -19,6 +19,11 @@ data DimConst = Zero | One
|
|||
|
||||
%runElab derive "DimConst" [Generic, Meta, Eq, Ord, DecEq, Show]
|
||||
|
||||
public export
|
||||
pick : a -> a -> DimConst -> a
|
||||
pick x y Zero = x
|
||||
pick x y One = y
|
||||
|
||||
|
||||
public export
|
||||
data Dim : Nat -> Type where
|
||||
|
|
|
@ -22,18 +22,27 @@ import Data.Vect
|
|||
%default total
|
||||
|
||||
|
||||
public export
|
||||
0 TermLike : Type
|
||||
TermLike = Type -> Nat -> Nat -> Type
|
||||
|
||||
public export
|
||||
0 TSubstLike : Type
|
||||
TSubstLike = Type -> Nat -> Nat -> Nat -> Type
|
||||
|
||||
|
||||
infixl 8 :#
|
||||
infixl 9 :@, :%
|
||||
mutual
|
||||
public export
|
||||
0 TSubst : Type -> Nat -> Nat -> Nat -> Type
|
||||
0 TSubst : TSubstLike
|
||||
TSubst q d = Subst $ Elim q d
|
||||
|
||||
||| first argument `q` is quantity type;
|
||||
||| second argument `d` is dimension scope size;
|
||||
||| third `n` is term scope size
|
||||
public export
|
||||
data Term : (q : Type) -> (d, n : Nat) -> Type where
|
||||
data Term : TermLike where
|
||||
||| type of types
|
||||
TYPE : (l : Universe) -> Term q d n
|
||||
|
||||
|
@ -61,7 +70,7 @@ mutual
|
|||
|
||||
||| first argument `d` is dimension scope size, second `n` is term scope size
|
||||
public export
|
||||
data Elim : (q : Type) -> (d, n : Nat) -> Type where
|
||||
data Elim : TermLike where
|
||||
||| free variable
|
||||
F : (x : Name) -> Elim q d n
|
||||
||| bound variable
|
||||
|
@ -85,7 +94,7 @@ mutual
|
|||
|
||||
||| a scope with one more bound variable
|
||||
public export
|
||||
data ScopeTerm : (q : Type) -> (d, n : Nat) -> Type where
|
||||
data ScopeTerm : TermLike where
|
||||
||| variable is used
|
||||
TUsed : (body : Term q d (S n)) -> ScopeTerm q d n
|
||||
||| variable is unused
|
||||
|
@ -93,7 +102,7 @@ mutual
|
|||
|
||||
||| a scope with one more bound dimension variable
|
||||
public export
|
||||
data DScopeTerm : (q : Type) -> (d, n : Nat) -> Type where
|
||||
data DScopeTerm : TermLike where
|
||||
||| variable is used
|
||||
DUsed : (body : Term q (S d) n) -> DScopeTerm q d n
|
||||
||| variable is unused
|
||||
|
|
|
@ -1,136 +1,121 @@
|
|||
module Quox.Syntax.Term.Reduce
|
||||
|
||||
import Quox.No
|
||||
import Quox.Syntax.Term.Base
|
||||
import Quox.Syntax.Term.Subst
|
||||
import Data.Maybe
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
mutual
|
||||
public export
|
||||
data NotCloT : Term {} -> Type where
|
||||
NCTYPE : NotCloT $ TYPE _
|
||||
NCPi : NotCloT $ Pi {}
|
||||
NCLam : NotCloT $ Lam {}
|
||||
NCEq : NotCloT $ Eq {}
|
||||
NCDLam : NotCloT $ DLam {}
|
||||
NCE : NotCloE e -> NotCloT $ E e
|
||||
namespace Elim
|
||||
public export %inline
|
||||
isClo : Elim {} -> Bool
|
||||
isClo (CloE {}) = True
|
||||
isClo (DCloE {}) = True
|
||||
isClo _ = False
|
||||
|
||||
public export
|
||||
data NotCloE : Elim {} -> Type where
|
||||
NCF : NotCloE $ F _
|
||||
NCB : NotCloE $ B _
|
||||
NCApp : NotCloE $ _ :@ _
|
||||
NCDApp : NotCloE $ _ :% _
|
||||
NCAnn : NotCloE $ _ :# _
|
||||
0 NotClo : Pred $ Elim {}
|
||||
NotClo = No . isClo
|
||||
|
||||
mutual
|
||||
export
|
||||
notCloT : (t : Term {}) -> Dec (NotCloT t)
|
||||
notCloT (TYPE _) = Yes NCTYPE
|
||||
notCloT (Pi {}) = Yes NCPi
|
||||
notCloT (Lam {}) = Yes NCLam
|
||||
notCloT (Eq {}) = Yes NCEq
|
||||
notCloT (DLam {}) = Yes NCDLam
|
||||
notCloT (E e) = case notCloE e of
|
||||
Yes nc => Yes $ NCE nc
|
||||
No c => No $ \case NCE nc => c nc
|
||||
notCloT (CloT {}) = No $ \case _ impossible
|
||||
notCloT (DCloT {}) = No $ \case _ impossible
|
||||
namespace Term
|
||||
public export %inline
|
||||
isClo : Term {} -> Bool
|
||||
isClo (CloT {}) = True
|
||||
isClo (DCloT {}) = True
|
||||
isClo (E e) = isClo e
|
||||
isClo _ = False
|
||||
|
||||
export
|
||||
notCloE : (e : Elim {}) -> Dec (NotCloE e)
|
||||
notCloE (F _) = Yes NCF
|
||||
notCloE (B _) = Yes NCB
|
||||
notCloE (_ :@ _) = Yes NCApp
|
||||
notCloE (_ :% _) = Yes NCDApp
|
||||
notCloE (_ :# _) = Yes NCAnn
|
||||
notCloE (CloE {}) = No $ \case _ impossible
|
||||
notCloE (DCloE {}) = No $ \case _ impossible
|
||||
public export
|
||||
0 NotClo : Pred $ Term {}
|
||||
NotClo = No . isClo
|
||||
|
||||
||| a term which is not a top level closure
|
||||
public export
|
||||
NonCloTerm : Type -> Nat -> Nat -> Type
|
||||
NonCloTerm q d n = Subset (Term q d n) NotCloT
|
||||
0 NonCloElim : TermLike
|
||||
NonCloElim q d n = Subset (Elim q d n) NotClo
|
||||
|
||||
||| an elimination which is not a top level closure
|
||||
public export
|
||||
NonCloElim : Type -> Nat -> Nat -> Type
|
||||
NonCloElim q d n = Subset (Elim q d n) NotCloE
|
||||
0 NonCloTerm : TermLike
|
||||
NonCloTerm q d n = Subset (Term q d n) NotClo
|
||||
|
||||
|
||||
public export %inline
|
||||
ncloT : (t : Term q d n) -> (0 _ : NotCloT t) => NonCloTerm q d n
|
||||
ncloT t @{p} = Element t p
|
||||
ncloT : (t : Term q d n) -> (0 nc : NotClo t) => NonCloTerm q d n
|
||||
ncloT t = Element t nc
|
||||
|
||||
public export %inline
|
||||
ncloE : (e : Elim q d n) -> (0 _ : NotCloE e) => NonCloElim q d n
|
||||
ncloE e @{p} = Element e p
|
||||
|
||||
ncloE : (e : Elim q d n) -> (0 nc : NotClo e) => NonCloElim q d n
|
||||
ncloE e = Element e nc
|
||||
|
||||
|
||||
mutual
|
||||
||| if the input term has any top-level closures, push them under one layer of
|
||||
||| syntax
|
||||
export %inline
|
||||
pushSubstsT : Term q d n -> NonCloTerm q d n
|
||||
pushSubstsT s = pushSubstsTWith id id s
|
||||
namespace Term
|
||||
||| if the input term has any top-level closures, push them under one layer of
|
||||
||| syntax
|
||||
export %inline
|
||||
pushSubsts : Term q d n -> NonCloTerm q d n
|
||||
pushSubsts s = pushSubstsWith id id s
|
||||
|
||||
||| if the input elimination has any top-level closures, push them under one
|
||||
||| layer of syntax
|
||||
export %inline
|
||||
pushSubstsE : Elim q d n -> NonCloElim q d n
|
||||
pushSubstsE e = pushSubstsEWith id id e
|
||||
export
|
||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
Term q dfrom from -> NonCloTerm q dto to
|
||||
pushSubstsWith th ph (TYPE l) =
|
||||
ncloT $ TYPE l
|
||||
pushSubstsWith th ph (Pi qty x a body) =
|
||||
ncloT $ Pi qty x (subs a th ph) (subs body th ph)
|
||||
pushSubstsWith th ph (Lam x body) =
|
||||
ncloT $ Lam x $ subs body th ph
|
||||
pushSubstsWith th ph (Eq i ty l r) =
|
||||
ncloT $ Eq i (subs ty th ph) (subs l th ph) (subs r th ph)
|
||||
pushSubstsWith th ph (DLam i body) =
|
||||
ncloT $ DLam i $ subs body th ph
|
||||
pushSubstsWith th ph (E e) =
|
||||
let Element e nc = pushSubstsWith th ph e in ncloT $ E e
|
||||
pushSubstsWith th ph (CloT s ps) =
|
||||
pushSubstsWith th (comp th ps ph) s
|
||||
pushSubstsWith th ph (DCloT s ps) =
|
||||
pushSubstsWith (ps . th) ph s
|
||||
|
||||
export
|
||||
pushSubstsTWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
Term q dfrom from -> NonCloTerm q dto to
|
||||
pushSubstsTWith th ph (TYPE l) =
|
||||
ncloT $ TYPE l
|
||||
pushSubstsTWith th ph (Pi qty x a body) =
|
||||
ncloT $ Pi qty x (subs a th ph) (subs body th ph)
|
||||
pushSubstsTWith th ph (Lam x body) =
|
||||
ncloT $ Lam x $ subs body th ph
|
||||
pushSubstsTWith th ph (Eq i ty l r) =
|
||||
ncloT $ Eq i (subs ty th ph) (subs l th ph) (subs r th ph)
|
||||
pushSubstsTWith th ph (DLam i body) =
|
||||
ncloT $ DLam i $ subs body th ph
|
||||
pushSubstsTWith th ph (E e) =
|
||||
let Element e nc = pushSubstsEWith th ph e in ncloT $ E e
|
||||
pushSubstsTWith th ph (CloT s ps) =
|
||||
pushSubstsTWith th (comp th ps ph) s
|
||||
pushSubstsTWith th ph (DCloT s ps) =
|
||||
pushSubstsTWith (ps . th) ph s
|
||||
namespace Elim
|
||||
||| if the input elimination has any top-level closures, push them under one
|
||||
||| layer of syntax
|
||||
export %inline
|
||||
pushSubsts : Elim q d n -> NonCloElim q d n
|
||||
pushSubsts e = pushSubstsWith id id e
|
||||
|
||||
export
|
||||
pushSubstsEWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
Elim q dfrom from -> NonCloElim q dto to
|
||||
pushSubstsEWith th ph (F x) =
|
||||
ncloE $ F x
|
||||
pushSubstsEWith th ph (B i) =
|
||||
let res = ph !! i in
|
||||
case notCloE res of
|
||||
Yes _ => ncloE res
|
||||
No _ => assert_total pushSubstsE res
|
||||
pushSubstsEWith th ph (f :@ s) =
|
||||
ncloE $ subs f th ph :@ subs s th ph
|
||||
pushSubstsEWith th ph (f :% d) =
|
||||
ncloE $ subs f th ph :% (d // th)
|
||||
pushSubstsEWith th ph (s :# a) =
|
||||
ncloE $ subs s th ph :# subs a th ph
|
||||
pushSubstsEWith th ph (CloE e ps) =
|
||||
pushSubstsEWith th (comp th ps ph) e
|
||||
pushSubstsEWith th ph (DCloE e ps) =
|
||||
pushSubstsEWith (ps . th) ph e
|
||||
export
|
||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
Elim q dfrom from -> NonCloElim q dto to
|
||||
pushSubstsWith th ph (F x) =
|
||||
ncloE $ F x
|
||||
pushSubstsWith th ph (B i) =
|
||||
let res = ph !! i in
|
||||
case nchoose $ isClo res of
|
||||
Left yes => assert_total pushSubsts res
|
||||
Right no => Element res no
|
||||
pushSubstsWith th ph (f :@ s) =
|
||||
ncloE $ subs f th ph :@ subs s th ph
|
||||
pushSubstsWith th ph (f :% d) =
|
||||
ncloE $ subs f th ph :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
ncloE $ subs s th ph :# subs a th ph
|
||||
pushSubstsWith th ph (CloE e ps) =
|
||||
pushSubstsWith th (comp th ps ph) e
|
||||
pushSubstsWith th ph (DCloE e ps) =
|
||||
pushSubstsWith (ps . th) ph e
|
||||
|
||||
|
||||
parameters (th : DSubst dfrom dto) (ph : TSubst q dto from to)
|
||||
public export %inline
|
||||
pushSubstsTWith' : Term q dfrom from -> Term q dto to
|
||||
pushSubstsTWith' s = (pushSubstsTWith th ph s).fst
|
||||
namespace Term
|
||||
public export %inline
|
||||
pushSubstsWith' : Term q dfrom from -> Term q dto to
|
||||
pushSubstsWith' s = (pushSubstsWith th ph s).fst
|
||||
|
||||
public export %inline
|
||||
pushSubstsEWith' : Elim q dfrom from -> Elim q dto to
|
||||
pushSubstsEWith' e = (pushSubstsEWith th ph e).fst
|
||||
namespace Elim
|
||||
public export %inline
|
||||
pushSubstsWith' : Elim q dfrom from -> Elim q dto to
|
||||
pushSubstsWith' e = (pushSubstsWith th ph e).fst
|
||||
|
||||
|
||||
public export %inline
|
||||
|
@ -142,197 +127,126 @@ weakE : Elim q d n -> Elim q d (S n)
|
|||
weakE t = t //. shift 1
|
||||
|
||||
|
||||
mutual
|
||||
public export
|
||||
data IsRedexT : Term q d n -> Type where
|
||||
IsUpsilonT : IsRedexT $ E (_ :# _)
|
||||
IsCloT : IsRedexT $ CloT {}
|
||||
IsDCloT : IsRedexT $ DCloT {}
|
||||
IsERedex : IsRedexE e -> IsRedexT $ E e
|
||||
|
||||
public export
|
||||
data IsRedexE : Elim q d n -> Type where
|
||||
IsUpsilonE : IsRedexE $ E _ :# _
|
||||
IsBetaLam : IsRedexE $ (Lam {} :# Pi {}) :@ _
|
||||
IsBetaDLam : IsRedexE $ (DLam {} :# Eq {}) :% _
|
||||
IsCloE : IsRedexE $ CloE {}
|
||||
IsDCloE : IsRedexE $ DCloE {}
|
||||
public export 0
|
||||
Lookup : TermLike
|
||||
Lookup q d n = Name -> Maybe $ Elim q d n
|
||||
|
||||
public export %inline
|
||||
NotRedexT : Term q d n -> Type
|
||||
NotRedexT = Not . IsRedexT
|
||||
isLamHead : Elim {} -> Bool
|
||||
isLamHead (Lam {} :# Pi {}) = True
|
||||
isLamHead _ = False
|
||||
|
||||
public export %inline
|
||||
NotRedexE : Elim q d n -> Type
|
||||
NotRedexE = Not . IsRedexE
|
||||
|
||||
|
||||
mutual
|
||||
-- [todo] PLEASE replace these with macros omfg
|
||||
export
|
||||
isRedexT : (t : Term {}) -> Dec (IsRedexT t)
|
||||
isRedexT (E (tm :# ty)) = Yes IsUpsilonT
|
||||
isRedexT (CloT {}) = Yes IsCloT
|
||||
isRedexT (DCloT {}) = Yes IsDCloT
|
||||
isRedexT (E (CloE {})) = Yes $ IsERedex IsCloE
|
||||
isRedexT (E (DCloE {})) = Yes $ IsERedex IsDCloE
|
||||
isRedexT (E e@(_ :@ _)) with (isRedexE e)
|
||||
_ | Yes yes = Yes $ IsERedex yes
|
||||
_ | No no = No $ \case IsERedex p => no p
|
||||
isRedexT (E e@(_ :% _)) with (isRedexE e)
|
||||
_ | Yes yes = Yes $ IsERedex yes
|
||||
_ | No no = No $ \case IsERedex p => no p
|
||||
isRedexT (TYPE {}) = No $ \case _ impossible
|
||||
isRedexT (Pi {}) = No $ \case _ impossible
|
||||
isRedexT (Lam {}) = No $ \case _ impossible
|
||||
isRedexT (Eq {}) = No $ \case _ impossible
|
||||
isRedexT (DLam {}) = No $ \case _ impossible
|
||||
isRedexT (E (F _)) = No $ \case IsERedex _ impossible
|
||||
isRedexT (E (B _)) = No $ \case IsERedex _ impossible
|
||||
|
||||
export
|
||||
isRedexE : (e : Elim {}) -> Dec (IsRedexE e)
|
||||
isRedexE (E _ :# _) = Yes IsUpsilonE
|
||||
isRedexE ((Lam {} :# Pi {}) :@ _) = Yes IsBetaLam
|
||||
isRedexE ((DLam {} :# Eq {}) :% _) = Yes IsBetaDLam
|
||||
isRedexE (CloE {}) = Yes IsCloE
|
||||
isRedexE (DCloE {}) = Yes IsDCloE
|
||||
isRedexE (F x) = No $ \case _ impossible
|
||||
isRedexE (B i) = No $ \case _ impossible
|
||||
isRedexE (F _ :@ _) = No $ \case _ impossible
|
||||
isRedexE (B _ :@ _) = No $ \case _ impossible
|
||||
isRedexE (_ :@ _ :@ _) = No $ \case _ impossible
|
||||
isRedexE (_ :% _ :@ _) = No $ \case _ impossible
|
||||
isRedexE (CloE {} :@ _) = No $ \case _ impossible
|
||||
isRedexE (DCloE {} :@ _) = No $ \case _ impossible
|
||||
isRedexE ((TYPE _ :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Pi {} :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Eq {} :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# TYPE _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# Lam {}) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# Eq {}) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# DLam {}) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# E _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# CloT {}) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# DCloT {}) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((E _ :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((CloT {} :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((DCloT {} :# _) :@ _) = No $ \case _ impossible
|
||||
isRedexE ((TYPE _ :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((Pi {} :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((Eq {} :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((Lam {} :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# TYPE _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# Pi {}) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# Lam {}) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# DLam {}) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# E _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# CloT {}) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DLam {} :# DCloT {}) :% _) = No $ \case _ impossible
|
||||
isRedexE ((E _ :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((CloT {} :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE ((DCloT {} :# _) :% _) = No $ \case _ impossible
|
||||
isRedexE (F _ :% _) = No $ \case _ impossible
|
||||
isRedexE (B _ :% _) = No $ \case _ impossible
|
||||
isRedexE (_ :@ _ :% _) = No $ \case _ impossible
|
||||
isRedexE (_ :% _ :% _) = No $ \case _ impossible
|
||||
isRedexE (CloE {} :% _) = No $ \case _ impossible
|
||||
isRedexE (DCloE {} :% _) = No $ \case _ impossible
|
||||
isRedexE (TYPE _ :# _) = No $ \case _ impossible
|
||||
isRedexE (Pi {} :# _) = No $ \case _ impossible
|
||||
isRedexE (Lam {} :# _) = No $ \case _ impossible
|
||||
isRedexE (Eq {} :# _) = No $ \case _ impossible
|
||||
isRedexE (DLam {} :# _) = No $ \case _ impossible
|
||||
isRedexE (CloT {} :# _) = No $ \case _ impossible
|
||||
isRedexE (DCloT {} :# _) = No $ \case _ impossible
|
||||
|
||||
isDLamHead : Elim {} -> Bool
|
||||
isDLamHead (DLam {} :# Eq {}) = True
|
||||
isDLamHead _ = False
|
||||
|
||||
public export %inline
|
||||
RedexTerm : Type -> Nat -> Nat -> Type
|
||||
RedexTerm q d n = Subset (Term q d n) IsRedexT
|
||||
isE : Term {} -> Bool
|
||||
isE (E _) = True
|
||||
isE _ = False
|
||||
|
||||
public export %inline
|
||||
NonRedexTerm : Type -> Nat -> Nat -> Type
|
||||
NonRedexTerm q d n = Subset (Term q d n) NotRedexT
|
||||
isAnn : Elim {} -> Bool
|
||||
isAnn (_ :# _) = True
|
||||
isAnn _ = False
|
||||
|
||||
public export %inline
|
||||
RedexElim : Type -> Nat -> Nat -> Type
|
||||
RedexElim q d n = Subset (Elim q d n) IsRedexE
|
||||
parameters (g : Lookup q d n)
|
||||
mutual
|
||||
namespace Elim
|
||||
public export
|
||||
isRedex : Elim q d n -> Bool
|
||||
isRedex (F x) = isJust $ g x
|
||||
isRedex (B _) = False
|
||||
isRedex (f :@ _) = isRedex f || isLamHead f
|
||||
isRedex (f :% _) = isRedex f || isDLamHead f
|
||||
isRedex (t :# a) = isE t || isRedex t || isRedex a
|
||||
isRedex (CloE {}) = True
|
||||
isRedex (DCloE {}) = True
|
||||
|
||||
public export %inline
|
||||
NonRedexElim : Type -> Nat -> Nat -> Type
|
||||
NonRedexElim q d n = Subset (Elim q d n) NotRedexE
|
||||
namespace Term
|
||||
public export
|
||||
isRedex : Term q d n -> Bool
|
||||
isRedex (CloT {}) = True
|
||||
isRedex (DCloT {}) = True
|
||||
isRedex (E e) = isAnn e || isRedex e
|
||||
isRedex _ = False
|
||||
|
||||
namespace Elim
|
||||
public export
|
||||
0 IsRedex, NotRedex : Pred $ Elim q d n
|
||||
IsRedex = So . isRedex
|
||||
NotRedex = No . isRedex
|
||||
|
||||
namespace Term
|
||||
public export
|
||||
0 IsRedex, NotRedex : Pred $ Term q d n
|
||||
IsRedex = So . isRedex
|
||||
NotRedex = No . isRedex
|
||||
|
||||
public export
|
||||
0 NonRedexElim, NonRedexTerm : (q, d, n : _) -> Lookup q d n -> Type
|
||||
NonRedexElim q d n g = Subset (Elim q d n) (NotRedex g)
|
||||
NonRedexTerm q d n g = Subset (Term q d n) (NotRedex g)
|
||||
|
||||
|
||||
||| substitute a term with annotation for the bound variable of a `ScopeTerm`
|
||||
export %inline
|
||||
substScope : (arg, argTy : Term q d n) -> (body : ScopeTerm q d n) -> Term q d n
|
||||
substScope arg argTy body = sub1 body (arg :# argTy)
|
||||
parameters (g : Lookup q d n)
|
||||
mutual
|
||||
namespace Elim
|
||||
export covering
|
||||
whnf : Elim q d n -> NonRedexElim q d n g
|
||||
whnf (F x) with (g x) proof eq
|
||||
_ | Just y = whnf y
|
||||
_ | Nothing = Element (F x) $ rewrite eq in Ah
|
||||
|
||||
mutual
|
||||
export %inline
|
||||
stepT' : (s : Term q d n) -> IsRedexT s -> Term q d n
|
||||
stepT' (E (s :# _)) IsUpsilonT = s
|
||||
stepT' (CloT s th) IsCloT = pushSubstsTWith' id th s
|
||||
stepT' (DCloT s th) IsDCloT = pushSubstsTWith' th id s
|
||||
stepT' (E e) (IsERedex p) = E $ stepE' e p
|
||||
whnf (B i) = Element (B i) Ah
|
||||
|
||||
export %inline
|
||||
stepE' : (e : Elim q d n) -> IsRedexE e -> Elim q d n
|
||||
stepE' (E e :# _) IsUpsilonE = e
|
||||
stepE' ((Lam {body, _} :# Pi {arg, res, _}) :@ s) IsBetaLam =
|
||||
let s = s :# arg in sub1 body s :# sub1 res s
|
||||
stepE' ((DLam {body, _} :# Eq {ty, l, r, _}) :% dim) IsBetaDLam =
|
||||
case dim of
|
||||
K Zero => l :# ty.zero
|
||||
K One => r :# ty.one
|
||||
B _ => dsub1 body dim :# dsub1 ty dim
|
||||
stepE' (CloE e th) IsCloE = pushSubstsEWith' id th e
|
||||
stepE' (DCloE e th) IsDCloE = pushSubstsEWith' th id e
|
||||
whnf (f :@ s) =
|
||||
let Element f fnf = whnf f in
|
||||
case nchoose $ isLamHead f of
|
||||
Left _ =>
|
||||
let Lam {body, _} :# Pi {arg, res, _} = f
|
||||
s = s :# arg
|
||||
in
|
||||
whnf $ sub1 body s :# sub1 res s
|
||||
Right nlh => Element (f :@ s) $ fnf `orNo` nlh
|
||||
|
||||
export %inline
|
||||
stepT : (s : Term q d n) -> Either (NotRedexT s) (Term q d n)
|
||||
stepT s = case isRedexT s of Yes y => Right $ stepT' s y; No n => Left n
|
||||
whnf (f :% p) =
|
||||
let Element f fnf = whnf f in
|
||||
case nchoose $ isDLamHead f of
|
||||
Left _ =>
|
||||
let DLam {body, _} :# Eq {ty, l, r, _} = f
|
||||
body = case p of K e => pick l r e; _ => dsub1 body p
|
||||
in
|
||||
whnf $ body :# dsub1 ty p
|
||||
Right ndlh =>
|
||||
Element (f :% p) $ fnf `orNo` ndlh
|
||||
|
||||
export %inline
|
||||
stepE : (e : Elim q d n) -> Either (NotRedexE e) (Elim q d n)
|
||||
stepE e = case isRedexE e of Yes y => Right $ stepE' e y; No n => Left n
|
||||
whnf (s :# a) =
|
||||
let Element s snf = whnf s
|
||||
Element a anf = whnf a
|
||||
in
|
||||
case nchoose $ isE s of
|
||||
Left _ => let E e = s in Element e $ noOr2 snf
|
||||
Right ne => Element (s :# a) $ ne `orNo` snf `orNo` anf
|
||||
|
||||
export covering
|
||||
whnfT : Term q d n -> NonRedexTerm q d n
|
||||
whnfT s = case stepT s of Right s' => whnfT s'; Left done => Element s done
|
||||
whnf (CloE el th) = whnf $ pushSubstsWith' id th el
|
||||
whnf (DCloE el th) = whnf $ pushSubstsWith' th id el
|
||||
|
||||
export covering
|
||||
whnfE : Elim q d n -> NonRedexElim q d n
|
||||
whnfE e = case stepE e of Right e' => whnfE e'; Left done => Element e done
|
||||
namespace Term
|
||||
export covering
|
||||
whnf : Term q d n -> NonRedexTerm q d n g
|
||||
whnf (TYPE l) = Element (TYPE l) Ah
|
||||
whnf (Pi qty x arg res) = Element (Pi qty x arg res) Ah
|
||||
whnf (Lam x body) = Element (Lam x body) Ah
|
||||
whnf (Eq i ty l r) = Element (Eq i ty l r) Ah
|
||||
whnf (DLam i body) = Element (DLam i body) Ah
|
||||
|
||||
whnf (E e) =
|
||||
let Element e enf = whnf e in
|
||||
case nchoose $ isAnn e of
|
||||
Left _ => let tm :# _ = e in Element tm $ noOr1 $ noOr2 enf
|
||||
Right na => Element (E e) $ na `orNo` enf
|
||||
|
||||
export
|
||||
notRedexNotCloE : (e : Elim {}) -> NotRedexE e -> NotCloE e
|
||||
notRedexNotCloE (F x) f = NCF
|
||||
notRedexNotCloE (B i) f = NCB
|
||||
notRedexNotCloE (fun :@ arg) f = NCApp
|
||||
notRedexNotCloE (fun :% arg) f = NCDApp
|
||||
notRedexNotCloE (tm :# ty) f = NCAnn
|
||||
notRedexNotCloE (CloE el th) f = absurd $ f IsCloE
|
||||
notRedexNotCloE (DCloE el th) f = absurd $ f IsDCloE
|
||||
|
||||
export
|
||||
notRedexNotCloT : (t : Term {}) -> NotRedexT t -> NotCloT t
|
||||
notRedexNotCloT (TYPE _) _ = NCTYPE
|
||||
notRedexNotCloT (Pi {}) _ = NCPi
|
||||
notRedexNotCloT (Lam {}) _ = NCLam
|
||||
notRedexNotCloT (Eq {}) _ = NCEq
|
||||
notRedexNotCloT (DLam {}) _ = NCDLam
|
||||
notRedexNotCloT (E e) f = NCE $ notRedexNotCloE e $ f . IsERedex
|
||||
notRedexNotCloT (CloT {}) f = absurd $ f IsCloT
|
||||
notRedexNotCloT (DCloT {}) f = absurd $ f IsDCloT
|
||||
|
||||
export
|
||||
toNotCloE : NonRedexElim q d n -> NonCloElim q d n
|
||||
toNotCloE (Element e prf) = Element e $ notRedexNotCloE e prf
|
||||
|
||||
export
|
||||
toNotCloT : NonRedexTerm q d n -> NonCloTerm q d n
|
||||
toNotCloT (Element t prf) = Element t $ notRedexNotCloT t prf
|
||||
whnf (CloT tm th) = whnf $ pushSubstsWith' id th tm
|
||||
whnf (DCloT tm th) = whnf $ pushSubstsWith' th id tm
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue