ScopeTerms that can bind multiple vars
This commit is contained in:
parent
92617a2e4a
commit
f0f49d9abf
5 changed files with 86 additions and 48 deletions
|
@ -165,13 +165,14 @@ parameters {0 isGlobal : _} (defs : Definitions' q isGlobal)
|
||||||
compare0 : CanEqual q m => Eq q => Elim q 0 n -> Elim q 0 n -> m ()
|
compare0 : CanEqual q m => Eq q => Elim q 0 n -> Elim q 0 n -> m ()
|
||||||
compare0 e f = compareN (whnf defs e) (whnf defs f)
|
compare0 e f = compareN (whnf defs e) (whnf defs f)
|
||||||
|
|
||||||
namespace ScopeTerm
|
namespace ScopeTermN
|
||||||
export covering %inline
|
export covering %inline
|
||||||
compare0 : CanEqual q m => Eq q =>
|
compare0 : {s : Nat} -> CanEqual q m => Eq q =>
|
||||||
ScopeTerm q 0 n -> ScopeTerm q 0 n -> m ()
|
ScopeTermN s q 0 n -> ScopeTermN s q 0 n -> m ()
|
||||||
compare0 (TUnused body0) (TUnused body1) = compare0 body0 body1
|
compare0 (TUnused body0) (TUnused body1) = compare0 body0 body1
|
||||||
compare0 body0 body1 = compare0 body0.term body1.term
|
compare0 body0 body1 = compare0 body0.term body1.term
|
||||||
|
|
||||||
|
-- [todo] extend to multi-var scopes
|
||||||
namespace DScopeTerm
|
namespace DScopeTerm
|
||||||
export covering %inline
|
export covering %inline
|
||||||
compare0 : CanEqual q m => Eq q =>
|
compare0 : CanEqual q m => Eq q =>
|
||||||
|
|
|
@ -5,6 +5,7 @@ import Quox.Syntax.Var
|
||||||
import Quox.Name
|
import Quox.Name
|
||||||
import Quox.Pretty
|
import Quox.Pretty
|
||||||
|
|
||||||
|
import Data.Nat
|
||||||
import Data.List
|
import Data.List
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
@ -84,6 +85,16 @@ public export %inline
|
||||||
push : CanSubst1 f => Subst f from to -> Subst f (S from) (S to)
|
push : CanSubst1 f => Subst f from to -> Subst f (S from) (S to)
|
||||||
push th = fromVar VZ ::: (th . shift 1)
|
push th = fromVar VZ ::: (th . shift 1)
|
||||||
|
|
||||||
|
-- [fixme] a better way to do this?
|
||||||
|
public export
|
||||||
|
pushN : CanSubst1 f => (s : Nat) ->
|
||||||
|
Subst f from to -> Subst f (s + from) (s + to)
|
||||||
|
pushN 0 th = th
|
||||||
|
pushN (S s) th =
|
||||||
|
rewrite plusSuccRightSucc s from in
|
||||||
|
rewrite plusSuccRightSucc s to in
|
||||||
|
pushN s $ fromVar VZ ::: (th . shift 1)
|
||||||
|
|
||||||
public export
|
public export
|
||||||
drop1 : Subst f (S from) to -> Subst f from to
|
drop1 : Subst f (S from) to -> Subst f from to
|
||||||
drop1 (Shift by) = Shift $ ssDown by
|
drop1 (Shift by) = Shift $ ssDown by
|
||||||
|
|
|
@ -92,26 +92,34 @@ mutual
|
||||||
DCloE : (el : Elim q dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
|
DCloE : (el : Elim q dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
|
||||||
Elim q dto n
|
Elim q dto n
|
||||||
|
|
||||||
||| a scope with one more bound variable
|
||| a scope with some more bound term variables
|
||||||
public export
|
public export
|
||||||
data ScopeTerm : TermLike where
|
data ScopeTermN : Nat -> TermLike where
|
||||||
||| variable is used
|
||| at least some variables are used
|
||||||
TUsed : (body : Term q d (S n)) -> ScopeTerm q d n
|
TUsed : (body : Term q d (s + n)) -> ScopeTermN s q d n
|
||||||
||| variable is unused
|
||| all variables are unused
|
||||||
TUnused : (body : Term q d n) -> ScopeTerm q d n
|
TUnused : (body : Term q d n) -> ScopeTermN s q d n
|
||||||
|
|
||||||
||| a scope with one more bound dimension variable
|
|
||||||
public export
|
public export
|
||||||
data DScopeTerm : TermLike where
|
0 ScopeTerm : TermLike
|
||||||
||| variable is used
|
ScopeTerm = ScopeTermN 1
|
||||||
DUsed : (body : Term q (S d) n) -> DScopeTerm q d n
|
|
||||||
||| variable is unused
|
||| a scope with some more bound dimension variable
|
||||||
DUnused : (body : Term q d n) -> DScopeTerm q d n
|
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
|
||||||
|
|
||||||
|
public export
|
||||||
|
0 DScopeTerm : TermLike
|
||||||
|
DScopeTerm = DScopeTermN 1
|
||||||
|
|
||||||
%name Term s, t, r
|
%name Term s, t, r
|
||||||
%name Elim e, f
|
%name Elim e, f
|
||||||
%name ScopeTerm body
|
%name ScopeTermN body
|
||||||
%name DScopeTerm body
|
%name DScopeTermN body
|
||||||
|
|
||||||
||| non dependent function type
|
||| non dependent function type
|
||||||
public export %inline
|
public export %inline
|
||||||
|
|
|
@ -82,7 +82,7 @@ mutual
|
||||||
[|withPrec SApp (prettyM e) </> prettyDSubst th|]
|
[|withPrec SApp (prettyM e) </> prettyDSubst th|]
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
PrettyHL q => PrettyHL (ScopeTerm q d n) where
|
{s : Nat} -> PrettyHL q => PrettyHL (ScopeTermN s q d n) where
|
||||||
prettyM body = prettyM body.term
|
prettyM body = prettyM body.term
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
|
|
|
@ -42,16 +42,18 @@ mutual
|
||||||
DCloE e ph /// th = DCloE e $ ph . th
|
DCloE e ph /// th = DCloE e $ ph . th
|
||||||
e /// th = DCloE e th
|
e /// th = DCloE e th
|
||||||
|
|
||||||
namespace ScopeTerm
|
namespace ScopeTermN
|
||||||
export
|
export
|
||||||
(///) : ScopeTerm q dfrom n -> DSubst dfrom dto -> ScopeTerm q dto n
|
(///) : ScopeTermN s q dfrom n -> DSubst dfrom dto -> ScopeTermN s q dto n
|
||||||
TUsed body /// th = TUsed $ body /// th
|
TUsed body /// th = TUsed $ body /// th
|
||||||
TUnused body /// th = TUnused $ body /// th
|
TUnused body /// th = TUnused $ body /// th
|
||||||
|
|
||||||
namespace DScopeTerm
|
namespace DScopeTermN
|
||||||
export
|
export
|
||||||
(///) : DScopeTerm q dfrom n -> DSubst dfrom dto -> DScopeTerm q dto n
|
(///) : {s : Nat} ->
|
||||||
DUsed body /// th = DUsed $ body /// push th
|
DScopeTermN s q dfrom n -> DSubst dfrom dto ->
|
||||||
|
DScopeTermN s q dto n
|
||||||
|
DUsed body /// th = DUsed $ body /// pushN s th
|
||||||
DUnused body /// th = DUnused $ body /// th
|
DUnused body /// th = DUnused $ body /// th
|
||||||
|
|
||||||
|
|
||||||
|
@ -89,19 +91,25 @@ CanSubst (Elim q d) (Term q d) where
|
||||||
th => CloT s th
|
th => CloT s th
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
CanSubst (Elim q d) (ScopeTerm q d) where
|
{s : Nat} -> CanSubst (Elim q d) (ScopeTermN s q d) where
|
||||||
TUsed body // th = TUsed $ body // push th
|
TUsed body // th = TUsed $ body // pushN s th
|
||||||
TUnused body // th = TUnused $ body // th
|
TUnused body // th = TUnused $ body // th
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
CanSubst (Elim q d) (DScopeTerm q d) where
|
{s : Nat} -> CanSubst (Elim q d) (DScopeTermN s q d) where
|
||||||
DUsed body // th = DUsed $ body // map (/// shift 1) th
|
DUsed body // th = DUsed $ body // map (/// shift s) th
|
||||||
DUnused body // th = DUnused $ body // th
|
DUnused body // th = DUnused $ body // th
|
||||||
|
|
||||||
export %inline CanSubst Var (Term q d) where s // th = s // map (B {q, d}) th
|
export %inline CanSubst Var (Term q d) where s // th = s // map (B {q, d}) th
|
||||||
export %inline CanSubst Var (Elim q d) where e // th = e // map (B {q, d}) th
|
export %inline CanSubst Var (Elim q d) where e // th = e // map (B {q, d}) th
|
||||||
export %inline CanSubst Var (ScopeTerm q d) where s // th = s // map (B {q, d}) th
|
|
||||||
export %inline CanSubst Var (DScopeTerm q d) where s // th = s // map (B {q, d}) th
|
export %inline
|
||||||
|
{s : Nat} -> CanSubst Var (ScopeTermN s q d) where
|
||||||
|
b // th = b // map (B {q, d}) th
|
||||||
|
|
||||||
|
export %inline
|
||||||
|
{s : Nat} -> CanSubst Var (DScopeTermN s q d) where
|
||||||
|
b // th = b // map (B {q, d}) th
|
||||||
|
|
||||||
|
|
||||||
infixl 8 //., ///
|
infixl 8 //., ///
|
||||||
|
@ -130,33 +138,43 @@ mutual
|
||||||
Elim q dto to
|
Elim q dto to
|
||||||
subs e th ph = e /// th // ph
|
subs e th ph = e /// th // ph
|
||||||
|
|
||||||
namespace ScopeTerm
|
namespace ScopeTermN
|
||||||
||| applies a term substitution with a less ambiguous type
|
||| applies a term substitution with a less ambiguous type
|
||||||
export %inline
|
export %inline
|
||||||
(//.) : ScopeTerm q d from -> TSubst q d from to -> ScopeTerm q d to
|
(//.) : {s : Nat} ->
|
||||||
|
ScopeTermN s q d from -> TSubst q d from to ->
|
||||||
|
ScopeTermN s q d to
|
||||||
body //. th = body // th
|
body //. th = body // th
|
||||||
|
|
||||||
||| applies a term and dimension substitution
|
||| applies a term and dimension substitution
|
||||||
public export %inline
|
public export %inline
|
||||||
subs : ScopeTerm q dfrom from -> DSubst dfrom dto -> TSubst q dto from to ->
|
subs : {s : Nat} ->
|
||||||
ScopeTerm q dto to
|
ScopeTermN s q dfrom from ->
|
||||||
|
DSubst dfrom dto -> TSubst q dto from to ->
|
||||||
|
ScopeTermN s q dto to
|
||||||
subs body th ph = body /// th // ph
|
subs body th ph = body /// th // ph
|
||||||
|
|
||||||
namespace DScopeTerm
|
namespace DScopeTermN
|
||||||
||| applies a term substitution with a less ambiguous type
|
||| applies a term substitution with a less ambiguous type
|
||||||
export %inline
|
export %inline
|
||||||
(//.) : DScopeTerm q d from -> TSubst q d from to -> DScopeTerm q d to
|
(//.) : {s : Nat} -> DScopeTermN s q d from -> TSubst q d from to ->
|
||||||
|
DScopeTermN s q d to
|
||||||
body //. th = body // th
|
body //. th = body // th
|
||||||
|
|
||||||
||| applies a term and dimension substitution
|
||| applies a term and dimension substitution
|
||||||
public export %inline
|
public export %inline
|
||||||
subs : DScopeTerm q dfrom from -> DSubst dfrom dto -> TSubst q dto from to ->
|
subs : {s : Nat} ->
|
||||||
DScopeTerm q dto to
|
DScopeTermN s q dfrom from ->
|
||||||
|
DSubst dfrom dto -> TSubst q dto from to ->
|
||||||
|
DScopeTermN s q dto to
|
||||||
subs body th ph = body /// th // ph
|
subs body th ph = body /// th // ph
|
||||||
|
|
||||||
export %inline CanShift (Term q d) where s // by = s //. Shift by
|
export %inline CanShift (Term q d) where s // by = s //. Shift by
|
||||||
export %inline CanShift (Elim q d) where e // by = e //. Shift by
|
export %inline CanShift (Elim q d) where e // by = e //. Shift by
|
||||||
export %inline CanShift (ScopeTerm q d) where s // by = s //. Shift by
|
|
||||||
|
export %inline
|
||||||
|
{s : Nat} -> CanShift (ScopeTermN s q d) where
|
||||||
|
b // by = b //. Shift by
|
||||||
|
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
|
@ -165,17 +183,17 @@ comp : DSubst dfrom dto -> TSubst q dfrom from mid -> TSubst q dto mid to ->
|
||||||
comp th ps ph = map (/// th) ps . ph
|
comp th ps ph = map (/// th) ps . ph
|
||||||
|
|
||||||
|
|
||||||
namespace ScopeTerm
|
namespace ScopeTermN
|
||||||
export %inline
|
export %inline
|
||||||
(.term) : ScopeTerm q d n -> Term q d (S n)
|
(.term) : {s : Nat} -> ScopeTermN s q d n -> Term q d (s + n)
|
||||||
(TUsed body).term = body
|
(TUsed body).term = body
|
||||||
(TUnused body).term = body //. shift 1
|
(TUnused body).term = body //. shift s
|
||||||
|
|
||||||
namespace DScopeTerm
|
namespace DScopeTermN
|
||||||
export %inline
|
export %inline
|
||||||
(.term) : DScopeTerm q d n -> Term q (S d) n
|
(.term) : {s : Nat} -> DScopeTermN s q d n -> Term q (s + d) n
|
||||||
(DUsed body).term = body
|
(DUsed body).term = body
|
||||||
(DUnused body).term = body /// shift 1
|
(DUnused body).term = body /// shift s
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
sub1 : ScopeTerm q d n -> Elim q d n -> Term q d n
|
sub1 : ScopeTerm q d n -> Elim q d n -> Term q d n
|
||||||
|
|
Loading…
Reference in a new issue