new representation for scopes
This commit is contained in:
parent
c75f1514ba
commit
0e481a8098
14 changed files with 376 additions and 364 deletions
|
@ -47,22 +47,20 @@ mutual
|
|||
TYPE : (l : Universe) -> Term q d n
|
||||
|
||||
||| function type
|
||||
Pi : (qty : q) -> (x : Name) ->
|
||||
(arg : Term q d n) -> (res : ScopeTerm q d n) -> Term q d n
|
||||
Pi : (qty : q) -> (arg : Term q d n) ->
|
||||
(res : ScopeTerm q d n) -> Term q d n
|
||||
||| function term
|
||||
Lam : (x : Name) -> (body : ScopeTerm q d n) -> Term q d n
|
||||
Lam : (body : ScopeTerm q d n) -> Term q d n
|
||||
|
||||
||| pair type
|
||||
Sig : (x : Name) -> (fst : Term q d n) -> (snd : ScopeTerm q d n) ->
|
||||
Term q d n
|
||||
Sig : (fst : Term q d n) -> (snd : ScopeTerm q d n) -> Term q d n
|
||||
||| pair value
|
||||
Pair : (fst, snd : Term q d n) -> Term q d n
|
||||
|
||||
||| equality type
|
||||
Eq : (i : Name) -> (ty : DScopeTerm q d n) -> (l, r : Term q d n) ->
|
||||
Term q d n
|
||||
Eq : (ty : DScopeTerm q d n) -> (l, r : Term q d n) -> Term q d n
|
||||
||| equality term
|
||||
DLam : (i : Name) -> (body : DScopeTerm q d n) -> Term q d n
|
||||
DLam : (body : DScopeTerm q d n) -> Term q d n
|
||||
|
||||
||| elimination
|
||||
E : (e : Elim q d n) -> Term q d n
|
||||
|
@ -87,11 +85,11 @@ mutual
|
|||
|
||||
||| pair destruction
|
||||
|||
|
||||
||| `CasePair 𝜋 𝑒 𝑟 𝐴 𝑥 𝑦 𝑡` is
|
||||
||| `CasePair 𝜋 𝑒 ([𝑟], 𝐴) ([𝑥, 𝑦], 𝑡)` is
|
||||
||| `𝐜𝐚𝐬𝐞 𝜋 · 𝑒 𝐫𝐞𝐭𝐮𝐫𝐧 𝑟 ⇒ 𝐴 𝐨𝐟 { (𝑥, 𝑦) ⇒ 𝑡 }`
|
||||
CasePair : (qty : q) -> (pair : Elim q d n) ->
|
||||
(r : Name) -> (ret : ScopeTerm q d n) ->
|
||||
(x, y : Name) -> (body : ScopeTermN 2 q d n) ->
|
||||
(ret : ScopeTerm q d n) ->
|
||||
(body : ScopeTermN 2 q d n) ->
|
||||
Elim q d n
|
||||
|
||||
||| dim application
|
||||
|
@ -107,49 +105,51 @@ mutual
|
|||
DCloE : (el : Elim q dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
|
||||
Elim q dto n
|
||||
|
||||
||| a scope with some more bound term variables
|
||||
public export
|
||||
data ScopeTermN : Nat -> TermLike where
|
||||
||| at least some variables are used
|
||||
TUsed : (body : Term q d (s + n)) -> ScopeTermN s q d n
|
||||
||| all variables are unused
|
||||
TUnused : (body : Term q d n) -> ScopeTermN s q d n
|
||||
0 CaseEnumArms : TermLike
|
||||
CaseEnumArms q d n = SortedMap TagVal (Term q d n)
|
||||
|
||||
||| a scoped term with names
|
||||
public export
|
||||
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
|
||||
constructor S
|
||||
names : Vect s BaseName
|
||||
body : ScopedBody s f n
|
||||
|
||||
public export
|
||||
0 ScopeTerm : TermLike
|
||||
ScopeTerm = ScopeTermN 1
|
||||
|
||||
||| a scope with some more bound dimension variable
|
||||
public export
|
||||
data DScopeTermN : Nat -> TermLike where
|
||||
||| at least some variables are used
|
||||
DUsed : (body : Term q (s + d) n) -> DScopeTermN s q d n
|
||||
||| all variables are unused
|
||||
DUnused : (body : Term q d n) -> DScopeTermN s q d n
|
||||
data ScopedBody : Nat -> (Nat -> Type) -> Nat -> Type where
|
||||
Y : (body : f (s + n)) -> ScopedBody s f n
|
||||
N : (body : f n) -> ScopedBody s f n
|
||||
|
||||
public export
|
||||
0 DScopeTerm : TermLike
|
||||
0 ScopeTermN, DScopeTermN : Nat -> TermLike
|
||||
ScopeTermN s q d n = Scoped s (Term q d) n
|
||||
DScopeTermN s q d n = Scoped s (\d => Term q d n) d
|
||||
|
||||
public export
|
||||
0 ScopeTerm, DScopeTerm : TermLike
|
||||
ScopeTerm = ScopeTermN 1
|
||||
DScopeTerm = DScopeTermN 1
|
||||
|
||||
%name Term s, t, r
|
||||
%name Elim e, f
|
||||
%name ScopeTermN body
|
||||
%name DScopeTermN body
|
||||
%name Scoped body
|
||||
%name ScopedBody body
|
||||
|
||||
||| non dependent function type
|
||||
public export %inline
|
||||
Arr : (qty : q) -> (arg, res : Term q d n) -> Term q d n
|
||||
Arr {qty, arg, res} = Pi {qty, x = "_", arg, res = TUnused res}
|
||||
Arr {qty, arg, res} = Pi {qty, arg, res = S ["_"] $ N res}
|
||||
|
||||
||| non dependent equality type
|
||||
public export %inline
|
||||
Eq0 : (ty, l, r : Term q d n) -> Term q d n
|
||||
Eq0 {ty, l, r} = Eq {i = "_", ty = DUnused ty, l, r}
|
||||
Eq0 {ty, l, r} = Eq {ty = S ["_"] $ N ty, l, r}
|
||||
|
||||
||| non dependent pair type
|
||||
public export %inline
|
||||
And : (fst, snd : Term q d n) -> Term q d n
|
||||
And {fst, snd} = Sig {x = "_", fst, snd = TUnused snd}
|
||||
And {fst, snd} = Sig {fst, snd = S ["_"] $ N snd}
|
||||
|
||||
||| same as `F` but as a term
|
||||
public export %inline
|
||||
|
|
|
@ -32,25 +32,80 @@ ofD = hl Syntax "of"
|
|||
returnD = hl Syntax "return"
|
||||
|
||||
|
||||
export covering
|
||||
prettyBindType : Pretty.HasEnv m =>
|
||||
PrettyHL q => PrettyHL a => PrettyHL b =>
|
||||
List q -> BaseName -> a -> Doc HL -> b -> m (Doc HL)
|
||||
prettyBindType qtys x s arr t =
|
||||
parensIfM Outer $ hang 2 $
|
||||
parens !(prettyQtyBinds qtys $ TV x) <++> arr
|
||||
<//> !(under T x $ prettyM t)
|
||||
|
||||
export covering
|
||||
prettyLams : Pretty.HasEnv m => PrettyHL a =>
|
||||
BinderSort -> List BaseName -> a -> m (Doc HL)
|
||||
prettyLams sort names body = do
|
||||
lam <- case sort of T => lamD; D => dlamD
|
||||
header <- sequence $ [hl TVar <$> prettyM x | x <- names] ++ [darrowD]
|
||||
body <- unders sort names $ prettyM body
|
||||
parensIfM Outer $ sep (lam :: header) <//> body
|
||||
|
||||
export covering
|
||||
prettyApps : Pretty.HasEnv m => PrettyHL f => PrettyHL a =>
|
||||
f -> List a -> m (Doc HL)
|
||||
prettyApps fun args =
|
||||
parensIfM App =<< withPrec Arg
|
||||
[|prettyM fun <//> (align . sep <$> traverse prettyM args)|]
|
||||
|
||||
export covering
|
||||
prettyTuple : Pretty.HasEnv m => PrettyHL a => List a -> m (Doc HL)
|
||||
prettyTuple = map (parens . align . separate commaD) . traverse prettyM
|
||||
|
||||
export covering
|
||||
prettyArm : Pretty.HasEnv m => PrettyHL a =>
|
||||
(List BaseName, Doc HL, a) -> m (Doc HL)
|
||||
prettyArm (xs, pat, body) =
|
||||
pure $ hang 2 $ sep
|
||||
[hsep [pat, !darrowD],
|
||||
!(withPrec Outer $ unders T xs $ prettyM body)]
|
||||
|
||||
export covering
|
||||
prettyArms : Pretty.HasEnv m => PrettyHL a =>
|
||||
List (List BaseName, Doc HL, a) -> m (Doc HL)
|
||||
prettyArms = map (braces . align . sep) . traverse prettyArm
|
||||
|
||||
export covering
|
||||
prettyCase : Pretty.HasEnv m =>
|
||||
PrettyHL q => PrettyHL a => PrettyHL b => PrettyHL c =>
|
||||
q -> a -> BaseName -> b -> List (List BaseName, Doc HL, c) ->
|
||||
m (Doc HL)
|
||||
prettyCase pi elim r ret arms =
|
||||
pure $ align $ sep $
|
||||
[hsep [caseD, !(prettyQtyBinds [pi] elim)],
|
||||
hsep [returnD, !(prettyM r), !darrowD, !(under T r $ prettyM ret)],
|
||||
hsep [ofD, !(prettyArms arms)]]
|
||||
|
||||
|
||||
mutual
|
||||
export covering
|
||||
PrettyHL q => PrettyHL (Term q d n) where
|
||||
prettyM (TYPE l) =
|
||||
parensIfM App $ typeD <//> !(withPrec Arg $ prettyM l)
|
||||
prettyM (Pi qty x s t) =
|
||||
prettyM (Pi qty s (S [x] t)) =
|
||||
prettyBindType [qty] x s !arrowD t
|
||||
prettyM (Lam x t) =
|
||||
let GotLams {names, body, _} = getLams' [x] t.term Refl in
|
||||
prettyM (Lam (S x t)) =
|
||||
let GotLams {names, body, _} = getLams' x t.term Refl in
|
||||
prettyLams T (toList names) body
|
||||
prettyM (Sig x s t) =
|
||||
prettyBindType [] x s !timesD t
|
||||
prettyM (Sig s (S [x] t)) =
|
||||
prettyBindType {q} [] x s !timesD t
|
||||
prettyM (Pair s t) =
|
||||
let GotPairs {init, last, _} = getPairs t in
|
||||
prettyTuple $ s :: init ++ [last]
|
||||
prettyM (Eq _ (DUnused ty) l r) =
|
||||
prettyM (Eq (S _ (N ty)) l r) =
|
||||
parensIfM Eq !(withPrec InEq $ pure $
|
||||
sep [!(prettyM l) <++> !eqndD, !(prettyM r) <++> colonD, !(prettyM ty)])
|
||||
prettyM (Eq i (DUsed ty) l r) =
|
||||
sep [!(prettyM l) <++> !eqndD,
|
||||
!(prettyM r) <++> colonD, !(prettyM ty)])
|
||||
prettyM (Eq (S [i] (Y ty)) l r) =
|
||||
parensIfM App $
|
||||
eqD <++>
|
||||
sep [bracks !(withPrec Outer $ pure $ hang 2 $
|
||||
|
@ -58,8 +113,8 @@ mutual
|
|||
!(under D i $ prettyM ty)]),
|
||||
!(withPrec Arg $ prettyM l),
|
||||
!(withPrec Arg $ prettyM r)]
|
||||
prettyM (DLam i t) =
|
||||
let GotDLams {names, body, _} = getDLams' [i] t.term Refl in
|
||||
prettyM (DLam (S i t)) =
|
||||
let GotDLams {names, body, _} = getDLams' i t.term Refl in
|
||||
prettyLams D (toList names) body
|
||||
prettyM (E e) = bracks <$> prettyM e
|
||||
prettyM (CloT s th) =
|
||||
|
@ -78,7 +133,7 @@ mutual
|
|||
prettyM (e :@ s) =
|
||||
let GotArgs {fun, args, _} = getArgs' e [s] in
|
||||
prettyApps fun args
|
||||
prettyM (CasePair pi p r ret x y body) = do
|
||||
prettyM (CasePair pi p (S [r] ret) (S [x, y] body)) = do
|
||||
pat <- (parens . separate commaD . map (hl TVar)) <$>
|
||||
traverse prettyM [x, y]
|
||||
prettyCase pi p r ret [([x, y], pat, body)]
|
||||
|
@ -97,62 +152,10 @@ mutual
|
|||
[|withPrec SApp (prettyM e) </> prettyDSubst th|]
|
||||
|
||||
export covering
|
||||
{s : Nat} -> PrettyHL q => PrettyHL (ScopeTermN s q d n) where
|
||||
{s : Nat} -> PrettyHL q => PrettyHL (ScopedBody s (Term q d) n) where
|
||||
prettyM body = prettyM body.term
|
||||
|
||||
export covering
|
||||
prettyTSubst : Pretty.HasEnv m => PrettyHL q =>
|
||||
TSubst q d from to -> m (Doc HL)
|
||||
prettyTSubst s = prettySubstM prettyM (!ask).tnames TVar "[" "]" s
|
||||
|
||||
export covering
|
||||
prettyBindType : Pretty.HasEnv m => PrettyHL q =>
|
||||
List q -> Name -> Term q d n -> Doc HL ->
|
||||
ScopeTerm q d n -> m (Doc HL)
|
||||
prettyBindType qtys x s arr t =
|
||||
parensIfM Outer $ hang 2 $
|
||||
parens !(prettyQtyBinds qtys $ TV x) <++> arr
|
||||
<//> !(under T x $ prettyM t)
|
||||
|
||||
export covering
|
||||
prettyLams : Pretty.HasEnv m => PrettyHL q =>
|
||||
BinderSort -> List Name -> Term q _ _ -> m (Doc HL)
|
||||
prettyLams sort names body = do
|
||||
lam <- case sort of T => lamD; D => dlamD
|
||||
header <- sequence $ [hl TVar <$> prettyM x | x <- names] ++ [darrowD]
|
||||
body <- unders sort names $ prettyM body
|
||||
parensIfM Outer $ sep (lam :: header) <//> body
|
||||
|
||||
export covering
|
||||
prettyApps : Pretty.HasEnv m => PrettyHL f => PrettyHL a =>
|
||||
f -> List a -> m (Doc HL)
|
||||
prettyApps fun args =
|
||||
parensIfM App =<< withPrec Arg
|
||||
[|prettyM fun <//> (align . sep <$> traverse prettyM args)|]
|
||||
|
||||
export covering
|
||||
prettyTuple : Pretty.HasEnv m => PrettyHL a => List a -> m (Doc HL)
|
||||
prettyTuple = map (parens . align . separate commaD) . traverse prettyM
|
||||
|
||||
export covering
|
||||
prettyArm : Pretty.HasEnv m => PrettyHL a =>
|
||||
(List Name, Doc HL, a) -> m (Doc HL)
|
||||
prettyArm (xs, pat, body) =
|
||||
pure $ hang 2 $ sep
|
||||
[hsep [pat, !darrowD],
|
||||
!(withPrec Outer $ unders T xs $ prettyM body)]
|
||||
|
||||
export covering
|
||||
prettyArms : Pretty.HasEnv m => PrettyHL a =>
|
||||
List (List Name, Doc HL, a) -> m (Doc HL)
|
||||
prettyArms = map (braces . align . sep) . traverse prettyArm
|
||||
|
||||
export covering
|
||||
prettyCase : Pretty.HasEnv m =>
|
||||
PrettyHL q => PrettyHL a => PrettyHL b => PrettyHL c =>
|
||||
q -> a -> Name -> b -> List (List Name, Doc HL, c) -> m (Doc HL)
|
||||
prettyCase pi elim r ret arms =
|
||||
pure $ align $ sep $
|
||||
[hsep [caseD, !(prettyQtyBinds [pi] elim)],
|
||||
hsep [returnD, !(prettyM r), !darrowD, !(under T r $ prettyM ret)],
|
||||
hsep [ofD, !(prettyArms arms)]]
|
||||
|
|
|
@ -11,8 +11,8 @@ import public Data.Vect
|
|||
|
||||
public export %inline
|
||||
isLam : Term {} -> Bool
|
||||
isLam (Lam {}) = True
|
||||
isLam _ = False
|
||||
isLam (Lam _) = True
|
||||
isLam _ = False
|
||||
|
||||
public export
|
||||
0 NotLam : Pred $ Term {}
|
||||
|
@ -21,8 +21,8 @@ NotLam = No . isLam
|
|||
|
||||
public export %inline
|
||||
isDLam : Term {} -> Bool
|
||||
isDLam (DLam {}) = True
|
||||
isDLam _ = False
|
||||
isDLam (DLam _) = True
|
||||
isDLam _ = False
|
||||
|
||||
public export
|
||||
0 NotDLam : Pred $ Term {}
|
||||
|
@ -113,25 +113,25 @@ getDArgs e = getDArgs' e []
|
|||
|
||||
infixr 1 :\\
|
||||
public export
|
||||
absN : Vect m Name -> Term q d (m + n) -> Term q d n
|
||||
absN : Vect m BaseName -> Term q d (m + n) -> Term q d n
|
||||
absN {m = 0} [] s = s
|
||||
absN {m = S m} (x :: xs) s = Lam x $ TUsed $ absN xs $
|
||||
rewrite sym $ plusSuccRightSucc m n in s
|
||||
absN {m = S m} (x :: xs) s =
|
||||
Lam $ S [x] $ Y $ absN xs $ rewrite sym $ plusSuccRightSucc m n in s
|
||||
|
||||
public export %inline
|
||||
(:\\) : Vect m Name -> Term q d (m + n) -> Term q d n
|
||||
(:\\) : Vect m BaseName -> Term q d (m + n) -> Term q d n
|
||||
(:\\) = absN
|
||||
|
||||
|
||||
infixr 1 :\\%
|
||||
public export
|
||||
dabsN : Vect m Name -> Term q (m + d) n -> Term q d n
|
||||
dabsN : Vect m BaseName -> Term q (m + d) n -> Term q d n
|
||||
dabsN {m = 0} [] s = s
|
||||
dabsN {m = S m} (x :: xs) s = DLam x $ DUsed $ dabsN xs $
|
||||
rewrite sym $ plusSuccRightSucc m d in s
|
||||
dabsN {m = S m} (x :: xs) s =
|
||||
DLam $ S [x] $ Y $ dabsN xs $ rewrite sym $ plusSuccRightSucc m d in s
|
||||
|
||||
public export %inline
|
||||
(:\\%) : Vect m Name -> Term q (m + d) n -> Term q d n
|
||||
(:\\%) : Vect m BaseName -> Term q (m + d) n -> Term q d n
|
||||
(:\\%) = dabsN
|
||||
|
||||
|
||||
|
@ -139,17 +139,17 @@ public export
|
|||
record GetLams q d n where
|
||||
constructor GotLams
|
||||
{0 lams, rest : Nat}
|
||||
names : Vect lams Name
|
||||
names : Vect lams BaseName
|
||||
body : Term q d rest
|
||||
0 eq : lams + n = rest
|
||||
0 notLam : NotLam body
|
||||
|
||||
export
|
||||
getLams' : forall lams, rest.
|
||||
Vect lams Name -> Term q d rest -> (0 eq : lams + n = rest) ->
|
||||
Vect lams BaseName -> Term q d rest -> (0 eq : lams + n = rest) ->
|
||||
GetLams q d n
|
||||
getLams' xs s Refl = case nchoose $ isLam s of
|
||||
Left y => let Lam x body = s in
|
||||
Left y => let Lam (S [x] body) = s in
|
||||
getLams' (x :: xs) (assert_smaller s body.term) Refl
|
||||
Right n => GotLams xs s Refl n
|
||||
|
||||
|
@ -162,17 +162,17 @@ public export
|
|||
record GetDLams q d n where
|
||||
constructor GotDLams
|
||||
{0 lams, rest : Nat}
|
||||
names : Vect lams Name
|
||||
names : Vect lams BaseName
|
||||
body : Term q rest n
|
||||
0 eq : lams + d = rest
|
||||
0 notDLam : NotDLam body
|
||||
|
||||
export
|
||||
getDLams' : forall lams, rest.
|
||||
Vect lams Name -> Term q rest n -> (0 eq : lams + d = rest) ->
|
||||
Vect lams BaseName -> Term q rest n -> (0 eq : lams + d = rest) ->
|
||||
GetDLams q d n
|
||||
getDLams' is s Refl = case nchoose $ isDLam s of
|
||||
Left y => let DLam i body = s in
|
||||
Left y => let DLam (S [i] body) = s in
|
||||
getDLams' (i :: is) (assert_smaller s body.term) Refl
|
||||
Right n => GotDLams is s Refl n
|
||||
|
||||
|
|
|
@ -10,49 +10,53 @@ namespace CanDSubst
|
|||
interface CanDSubst (0 tm : Nat -> Nat -> Type) where
|
||||
(//) : tm dfrom n -> Lazy (DSubst dfrom dto) -> tm dto n
|
||||
|
||||
mutual
|
||||
||| does the minimal reasonable work:
|
||||
||| - deletes the closure around an atomic constant like `TYPE`
|
||||
||| - deletes an identity substitution
|
||||
||| - composes (lazily) with an existing top-level dim-closure
|
||||
||| - otherwise, wraps in a new closure
|
||||
export
|
||||
CanDSubst (Term q) where
|
||||
s // Shift SZ = s
|
||||
TYPE l // _ = TYPE l
|
||||
DCloT s ph // th = DCloT s $ ph . th
|
||||
s // th = DCloT s th
|
||||
||| does the minimal reasonable work:
|
||||
||| - deletes the closure around an atomic constant like `TYPE`
|
||||
||| - deletes an identity substitution
|
||||
||| - composes (lazily) with an existing top-level dim-closure
|
||||
||| - otherwise, wraps in a new closure
|
||||
export
|
||||
CanDSubst (Term q) where
|
||||
s // Shift SZ = s
|
||||
TYPE l // _ = TYPE l
|
||||
DCloT s ph // th = DCloT s $ ph . th
|
||||
s // th = DCloT s th
|
||||
|
||||
private
|
||||
subDArgs : Elim q dfrom n -> DSubst dfrom dto -> Elim q dto n
|
||||
subDArgs (f :% d) th = subDArgs f th :% (d // th)
|
||||
subDArgs e th = DCloE e th
|
||||
private
|
||||
subDArgs : Elim q dfrom n -> DSubst dfrom dto -> Elim q dto n
|
||||
subDArgs (f :% d) th = subDArgs f th :% (d // th)
|
||||
subDArgs e th = DCloE e th
|
||||
|
||||
||| does the minimal reasonable work:
|
||||
||| - deletes the closure around a term variable
|
||||
||| - deletes an identity substitution
|
||||
||| - composes (lazily) with an existing top-level dim-closure
|
||||
||| - immediately looks up bound variables in a
|
||||
||| top-level sequence of dimension applications
|
||||
||| - otherwise, wraps in a new closure
|
||||
export
|
||||
CanDSubst (Elim q) where
|
||||
e // Shift SZ = e
|
||||
F x // _ = F x
|
||||
B i // _ = B i
|
||||
f :% d // th = subDArgs (f :% d) th
|
||||
DCloE e ph // th = DCloE e $ ph . th
|
||||
e // th = DCloE e th
|
||||
||| does the minimal reasonable work:
|
||||
||| - deletes the closure around a term variable
|
||||
||| - deletes an identity substitution
|
||||
||| - composes (lazily) with an existing top-level dim-closure
|
||||
||| - immediately looks up bound variables in a
|
||||
||| top-level sequence of dimension applications
|
||||
||| - otherwise, wraps in a new closure
|
||||
export
|
||||
CanDSubst (Elim q) where
|
||||
e // Shift SZ = e
|
||||
F x // _ = F x
|
||||
B i // _ = B i
|
||||
f :% d // th = subDArgs (f :% d) th
|
||||
DCloE e ph // th = DCloE e $ ph . th
|
||||
e // th = DCloE e th
|
||||
|
||||
export
|
||||
CanDSubst (ScopeTermN s q) where
|
||||
TUsed body // th = TUsed $ body // th
|
||||
TUnused body // th = TUnused $ body // th
|
||||
namespace DSubst.ScopeTermN
|
||||
export %inline
|
||||
(//) : ScopeTermN s q dfrom n -> Lazy (DSubst dfrom dto) ->
|
||||
ScopeTermN s q dto n
|
||||
S ns (Y body) // th = S ns $ Y $ body // th
|
||||
S ns (N body) // th = S ns $ N $ body // th
|
||||
|
||||
export
|
||||
{s : Nat} -> CanDSubst (DScopeTermN s q) where
|
||||
DUsed body // th = DUsed $ body // pushN s th
|
||||
DUnused body // th = DUnused $ body // th
|
||||
namespace DSubst.DScopeTermN
|
||||
export %inline
|
||||
(//) : {s : Nat} ->
|
||||
DScopeTermN s q dfrom n -> Lazy (DSubst dfrom dto) ->
|
||||
DScopeTermN s q dto n
|
||||
S ns (Y body) // th = S ns $ Y $ body // pushN s th
|
||||
S ns (N body) // th = S ns $ N $ body // th
|
||||
|
||||
|
||||
export %inline FromVar (Elim q d) where fromVar = B
|
||||
|
@ -94,15 +98,21 @@ CanTSubst Term where
|
|||
Shift SZ => s
|
||||
th => CloT s th
|
||||
|
||||
export %inline
|
||||
{s : Nat} -> CanTSubst (ScopeTermN s) where
|
||||
TUsed body // th = TUsed $ body // pushN s th
|
||||
TUnused body // th = TUnused $ body // th
|
||||
namespace ScopeTermN
|
||||
export %inline
|
||||
(//) : {s : Nat} ->
|
||||
ScopeTermN s q d from -> Lazy (TSubst q d from to) ->
|
||||
ScopeTermN s q d to
|
||||
S ns (Y body) // th = S ns $ Y $ body // pushN s th
|
||||
S ns (N body) // th = S ns $ N $ body // th
|
||||
|
||||
export %inline
|
||||
{s : Nat} -> CanTSubst (DScopeTermN s) where
|
||||
DUsed body // th = DUsed $ body // map (// shift s) th
|
||||
DUnused body // th = DUnused $ body // th
|
||||
namespace DScopeTermN
|
||||
export %inline
|
||||
(//) : {s : Nat} ->
|
||||
DScopeTermN s q d from -> Lazy (TSubst q d from to) ->
|
||||
DScopeTermN s q d to
|
||||
S ns (Y body) // th = S ns $ Y $ body // map (// shift s) th
|
||||
S ns (N body) // th = S ns $ N $ body // th
|
||||
|
||||
export %inline CanShift (Term q d) where s // by = s // Shift by
|
||||
export %inline CanShift (Elim q d) where e // by = e // Shift by
|
||||
|
@ -136,23 +146,34 @@ weakE : {default 1 by : Nat} -> Elim q d n -> Elim q d (by + n)
|
|||
weakE t = t // shift by
|
||||
|
||||
|
||||
namespace ScopeTermN
|
||||
export %inline
|
||||
(.term) : {s : Nat} -> ScopeTermN s q d n -> Term q d (s + n)
|
||||
(TUsed body).term = body
|
||||
(TUnused body).term = weakT body {by = s}
|
||||
parameters {s : Nat}
|
||||
namespace ScopeTermBody
|
||||
export %inline
|
||||
(.term) : ScopedBody s (Term q d) n -> Term q d (s + n)
|
||||
(Y b).term = b
|
||||
(N b).term = weakT b {by = s}
|
||||
|
||||
namespace DScopeTermN
|
||||
export %inline
|
||||
(.term) : {s : Nat} -> DScopeTermN s q d n -> Term q (s + d) n
|
||||
(DUsed body).term = body
|
||||
(DUnused body).term = dweakT body {by = s}
|
||||
namespace ScopeTermN
|
||||
export %inline
|
||||
(.term) : ScopeTermN s q d n -> Term q d (s + n)
|
||||
t.term = t.body.term
|
||||
|
||||
namespace DScopeTermBody
|
||||
export %inline
|
||||
(.term) : ScopedBody s (\d => Term q d n) d -> Term q (s + d) n
|
||||
(Y b).term = b
|
||||
(N b).term = dweakT b {by = s}
|
||||
|
||||
namespace DScopeTermN
|
||||
export %inline
|
||||
(.term) : DScopeTermN s q d n -> Term q (s + d) n
|
||||
t.term = t.body.term
|
||||
|
||||
|
||||
export %inline
|
||||
subN : ScopeTermN s q d n -> Vect s (Elim q d n) -> Term q d n
|
||||
subN (TUsed body) es = body // fromVect es
|
||||
subN (TUnused body) _ = body
|
||||
subN (S _ (Y body)) es = body // fromVect es
|
||||
subN (S _ (N body)) _ = body
|
||||
|
||||
export %inline
|
||||
sub1 : ScopeTerm q d n -> Elim q d n -> Term q d n
|
||||
|
@ -160,8 +181,8 @@ sub1 t e = subN t [e]
|
|||
|
||||
export %inline
|
||||
dsubN : DScopeTermN s q d n -> Vect s (Dim d) -> Term q d n
|
||||
dsubN (DUsed body) ps = body // fromVect ps
|
||||
dsubN (DUnused body) _ = body
|
||||
dsubN (S _ (Y body)) ps = body // fromVect ps
|
||||
dsubN (S _ (N body)) _ = body
|
||||
|
||||
export %inline
|
||||
dsub1 : DScopeTerm q d n -> Dim d -> Term q d n
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue