move pushSubsts to Q.S.T.Subst
This commit is contained in:
parent
55cdb19a4c
commit
60f07a938e
2 changed files with 103 additions and 104 deletions
|
@ -10,110 +10,6 @@ import Data.List
|
|||
%default total
|
||||
|
||||
|
||||
|
||||
public export
|
||||
0 CloTest : TermLike -> Type
|
||||
CloTest tm = forall q, d, n. tm q d n -> Bool
|
||||
|
||||
interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where
|
||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
tm q dfrom from -> Subset (tm q dto to) (No . isClo)
|
||||
|
||||
public export
|
||||
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm q d n)
|
||||
NotClo = No . isClo
|
||||
|
||||
public export
|
||||
0 NonClo : (tm : TermLike) -> {isClo : CloTest tm} ->
|
||||
PushSubsts tm isClo => TermLike
|
||||
NonClo tm q d n = Subset (tm q d n) NotClo
|
||||
|
||||
public export %inline
|
||||
nclo : {isClo : CloTest tm} -> (0 _ : PushSubsts tm isClo) =>
|
||||
(t : tm q d n) -> (0 nc : NotClo t) => NonClo tm q d n
|
||||
nclo t = Element t nc
|
||||
|
||||
parameters {0 isClo : CloTest tm} {auto _ : PushSubsts tm isClo}
|
||||
||| if the input term has any top-level closures, push them under one layer of
|
||||
||| syntax
|
||||
export %inline
|
||||
pushSubsts : tm q d n -> NonClo tm q d n
|
||||
pushSubsts s = pushSubstsWith id id s
|
||||
|
||||
export %inline
|
||||
pushSubstsWith' : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
tm q dfrom from -> tm q dto to
|
||||
pushSubstsWith' th ph x = fst $ pushSubstsWith th ph x
|
||||
|
||||
mutual
|
||||
public export
|
||||
isCloT : CloTest Term
|
||||
isCloT (CloT {}) = True
|
||||
isCloT (DCloT {}) = True
|
||||
isCloT (E e) = isCloE e
|
||||
isCloT _ = False
|
||||
|
||||
public export
|
||||
isCloE : CloTest Elim
|
||||
isCloE (CloE {}) = True
|
||||
isCloE (DCloE {}) = True
|
||||
isCloE _ = False
|
||||
|
||||
mutual
|
||||
export
|
||||
PushSubsts Term Reduce.isCloT where
|
||||
pushSubstsWith th ph (TYPE l) =
|
||||
nclo $ TYPE l
|
||||
pushSubstsWith th ph (Pi qty a body) =
|
||||
nclo $ Pi qty (a // th // ph) (body // th // ph)
|
||||
pushSubstsWith th ph (Lam body) =
|
||||
nclo $ Lam (body // th // ph)
|
||||
pushSubstsWith th ph (Sig a b) =
|
||||
nclo $ Sig (a // th // ph) (b // th // ph)
|
||||
pushSubstsWith th ph (Pair s t) =
|
||||
nclo $ Pair (s // th // ph) (t // th // ph)
|
||||
pushSubstsWith th ph (Enum tags) =
|
||||
nclo $ Enum tags
|
||||
pushSubstsWith th ph (Tag tag) =
|
||||
nclo $ Tag tag
|
||||
pushSubstsWith th ph (Eq ty l r) =
|
||||
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph)
|
||||
pushSubstsWith th ph (DLam body) =
|
||||
nclo $ DLam (body // th // ph)
|
||||
pushSubstsWith th ph (E e) =
|
||||
let Element e nc = pushSubstsWith th ph e in nclo $ 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
|
||||
PushSubsts Elim Reduce.isCloE where
|
||||
pushSubstsWith th ph (F x) =
|
||||
nclo $ F x
|
||||
pushSubstsWith th ph (B i) =
|
||||
let res = ph !! i in
|
||||
case nchoose $ isCloE res of
|
||||
Left yes => assert_total pushSubsts res
|
||||
Right no => Element res no
|
||||
pushSubstsWith th ph (f :@ s) =
|
||||
nclo $ (f // th // ph) :@ (s // th // ph)
|
||||
pushSubstsWith th ph (CasePair pi p r b) =
|
||||
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph)
|
||||
pushSubstsWith th ph (CaseEnum pi t r arms) =
|
||||
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
|
||||
(map (\b => b // th // ph) arms)
|
||||
pushSubstsWith th ph (f :% d) =
|
||||
nclo $ (f // th // ph) :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
nclo $ (s // th // ph) :# (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
|
||||
|
||||
|
||||
|
||||
||| errors that might happen if you pass an ill typed expression into
|
||||
||| whnf. don't do that please
|
||||
public export
|
||||
|
|
|
@ -1,5 +1,6 @@
|
|||
module Quox.Syntax.Term.Subst
|
||||
|
||||
import Quox.No
|
||||
import Quox.Syntax.Term.Base
|
||||
import Data.Vect
|
||||
|
||||
|
@ -196,3 +197,105 @@ body.zero = dsub1 body $ K Zero
|
|||
public export %inline
|
||||
(.one) : DScopeTerm q d n -> Term q d n
|
||||
body.one = dsub1 body $ K One
|
||||
|
||||
|
||||
public export
|
||||
0 CloTest : TermLike -> Type
|
||||
CloTest tm = forall q, d, n. tm q d n -> Bool
|
||||
|
||||
interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where
|
||||
pushSubstsWith : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
tm q dfrom from -> Subset (tm q dto to) (No . isClo)
|
||||
|
||||
public export
|
||||
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm q d n)
|
||||
NotClo = No . isClo
|
||||
|
||||
public export
|
||||
0 NonClo : (tm : TermLike) -> {isClo : CloTest tm} ->
|
||||
PushSubsts tm isClo => TermLike
|
||||
NonClo tm q d n = Subset (tm q d n) NotClo
|
||||
|
||||
public export %inline
|
||||
nclo : {isClo : CloTest tm} -> (0 _ : PushSubsts tm isClo) =>
|
||||
(t : tm q d n) -> (0 nc : NotClo t) => NonClo tm q d n
|
||||
nclo t = Element t nc
|
||||
|
||||
parameters {0 isClo : CloTest tm} {auto _ : PushSubsts tm isClo}
|
||||
||| if the input term has any top-level closures, push them under one layer of
|
||||
||| syntax
|
||||
export %inline
|
||||
pushSubsts : tm q d n -> NonClo tm q d n
|
||||
pushSubsts s = pushSubstsWith id id s
|
||||
|
||||
export %inline
|
||||
pushSubstsWith' : DSubst dfrom dto -> TSubst q dto from to ->
|
||||
tm q dfrom from -> tm q dto to
|
||||
pushSubstsWith' th ph x = fst $ pushSubstsWith th ph x
|
||||
|
||||
mutual
|
||||
public export
|
||||
isCloT : CloTest Term
|
||||
isCloT (CloT {}) = True
|
||||
isCloT (DCloT {}) = True
|
||||
isCloT (E e) = isCloE e
|
||||
isCloT _ = False
|
||||
|
||||
public export
|
||||
isCloE : CloTest Elim
|
||||
isCloE (CloE {}) = True
|
||||
isCloE (DCloE {}) = True
|
||||
isCloE _ = False
|
||||
|
||||
mutual
|
||||
export
|
||||
PushSubsts Term Subst.isCloT where
|
||||
pushSubstsWith th ph (TYPE l) =
|
||||
nclo $ TYPE l
|
||||
pushSubstsWith th ph (Pi qty a body) =
|
||||
nclo $ Pi qty (a // th // ph) (body // th // ph)
|
||||
pushSubstsWith th ph (Lam body) =
|
||||
nclo $ Lam (body // th // ph)
|
||||
pushSubstsWith th ph (Sig a b) =
|
||||
nclo $ Sig (a // th // ph) (b // th // ph)
|
||||
pushSubstsWith th ph (Pair s t) =
|
||||
nclo $ Pair (s // th // ph) (t // th // ph)
|
||||
pushSubstsWith th ph (Enum tags) =
|
||||
nclo $ Enum tags
|
||||
pushSubstsWith th ph (Tag tag) =
|
||||
nclo $ Tag tag
|
||||
pushSubstsWith th ph (Eq ty l r) =
|
||||
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph)
|
||||
pushSubstsWith th ph (DLam body) =
|
||||
nclo $ DLam (body // th // ph)
|
||||
pushSubstsWith th ph (E e) =
|
||||
let Element e nc = pushSubstsWith th ph e in nclo $ 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
|
||||
PushSubsts Elim Subst.isCloE where
|
||||
pushSubstsWith th ph (F x) =
|
||||
nclo $ F x
|
||||
pushSubstsWith th ph (B i) =
|
||||
let res = ph !! i in
|
||||
case nchoose $ isCloE res of
|
||||
Left yes => assert_total pushSubsts res
|
||||
Right no => Element res no
|
||||
pushSubstsWith th ph (f :@ s) =
|
||||
nclo $ (f // th // ph) :@ (s // th // ph)
|
||||
pushSubstsWith th ph (CasePair pi p r b) =
|
||||
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph)
|
||||
pushSubstsWith th ph (CaseEnum pi t r arms) =
|
||||
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
|
||||
(map (\b => b // th // ph) arms)
|
||||
pushSubstsWith th ph (f :% d) =
|
||||
nclo $ (f // th // ph) :% (d // th)
|
||||
pushSubstsWith th ph (s :# a) =
|
||||
nclo $ (s // th // ph) :# (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
|
||||
|
|
Loading…
Reference in a new issue