quox/lib/Quox/Syntax/Term/Subst.idr

339 lines
11 KiB
Idris
Raw Normal View History

2022-04-23 18:21:30 -04:00
module Quox.Syntax.Term.Subst
2023-02-26 05:17:42 -05:00
import Quox.No
2022-04-23 18:21:30 -04:00
import Quox.Syntax.Term.Base
2023-03-26 10:09:47 -04:00
import Data.SnocVect
2022-04-23 18:21:30 -04:00
2022-05-02 16:38:37 -04:00
%default total
namespace CanDSubst
public export
2023-04-01 13:16:43 -04:00
interface CanDSubst (0 tm : TermLike) where
2023-06-23 12:32:05 -04:00
(//) : tm d1 n -> Lazy (DSubst d1 d2) -> tm d2 n
2022-04-23 18:21:30 -04:00
2023-02-22 01:40:19 -05:00
||| 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
2023-04-01 13:16:43 -04:00
CanDSubst Term where
s // Shift SZ = s
2023-05-01 21:06:25 -04:00
TYPE l loc // _ = TYPE l loc
DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th
s // th = DCloT $ Sub s th
2023-02-22 01:40:19 -05:00
private
2023-06-23 12:32:05 -04:00
subDArgs : Elim d1 n -> DSubst d1 d2 -> Elim d2 n
2023-05-01 21:06:25 -04:00
subDArgs (DApp f d loc) th = DApp (subDArgs f th) (d // th) loc
subDArgs e th = DCloE $ Sub e th
2023-02-22 01:40:19 -05:00
||| 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
2023-04-01 13:16:43 -04:00
CanDSubst Elim where
e // Shift SZ = e
2023-05-21 14:09:34 -04:00
F x u loc // _ = F x u loc
2023-05-01 21:06:25 -04:00
B i loc // _ = B i loc
e@(DApp {}) // th = subDArgs e th
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
e // th = DCloE $ Sub e th
2023-02-22 01:40:19 -05:00
namespace DSubst.ScopeTermN
export %inline
2023-06-23 12:32:05 -04:00
(//) : ScopeTermN s d1 n -> Lazy (DSubst d1 d2) ->
ScopeTermN s d2 n
2023-02-22 01:40:19 -05:00
S ns (Y body) // th = S ns $ Y $ body // th
S ns (N body) // th = S ns $ N $ body // th
namespace DSubst.DScopeTermN
export %inline
(//) : {s : Nat} ->
2023-06-23 12:32:05 -04:00
DScopeTermN s d1 n -> Lazy (DSubst d1 d2) ->
DScopeTermN s d2 n
2023-11-27 15:01:36 -05:00
S ns (Y body) // th = S ns $ Y $ body // pushN s (locs $ toList' ns) th
2023-02-22 01:40:19 -05:00
S ns (N body) // th = S ns $ N $ body // th
2023-01-20 20:34:28 -05:00
2023-05-01 21:06:25 -04:00
export %inline FromVar (Elim d) where fromVarLoc = B
2023-11-27 15:01:36 -05:00
export %inline FromVar (Term d) where fromVarLoc = E .: fromVarLoc
2022-04-23 18:21:30 -04:00
2022-04-23 18:21:30 -04:00
||| does the minimal reasonable work:
2023-01-20 20:34:28 -05:00
||| - deletes the closure around a *free* name
2022-04-23 18:21:30 -04:00
||| - deletes an identity substitution
||| - composes (lazily) with an existing top-level closure
||| - immediately looks up a bound variable
||| - otherwise, wraps in a new closure
export
2023-04-01 13:16:43 -04:00
CanSubstSelf (Elim d) where
2023-05-21 14:09:34 -04:00
F x u loc // _ = F x u loc
2023-05-01 21:06:25 -04:00
B i loc // th = getLoc th i loc
CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th
e // th = case force th of
Shift SZ => e
th => CloE $ Sub e th
2022-04-23 18:21:30 -04:00
namespace CanTSubst
public export
interface CanTSubst (0 tm : TermLike) where
2023-06-23 12:32:05 -04:00
(//) : tm d n1 -> Lazy (TSubst d n1 n2) -> tm d n2
2022-04-23 18:21:30 -04:00
||| 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 closure
||| - goes inside `E` in case it is a simple variable or something
||| - otherwise, wraps in a new closure
export
CanTSubst Term where
2023-05-01 21:06:25 -04:00
TYPE l loc // _ = TYPE l loc
E e // th = E $ e // th
CloT (Sub s ph) // th = CloT $ Sub s $ ph . th
s // th = case force th of
Shift SZ => s
th => CloT $ Sub s th
2022-04-23 18:21:30 -04:00
2023-02-22 01:40:19 -05:00
namespace ScopeTermN
export %inline
(//) : {s : Nat} ->
2023-06-23 12:32:05 -04:00
ScopeTermN s d n1 -> Lazy (TSubst d n1 n2) ->
ScopeTermN s d n2
2023-11-27 15:01:36 -05:00
S ns (Y body) // th = S ns $ Y $ body // pushN s (locs $ toList' ns) th
2023-02-22 01:40:19 -05:00
S ns (N body) // th = S ns $ N $ body // th
2022-04-23 18:21:30 -04:00
2023-02-22 01:40:19 -05:00
namespace DScopeTermN
export %inline
(//) : {s : Nat} ->
2023-06-23 12:32:05 -04:00
DScopeTermN s d n1 -> Lazy (TSubst d n1 n2) -> DScopeTermN s d n2
2023-02-22 01:40:19 -05:00
S ns (Y body) // th = S ns $ Y $ body // map (// shift s) th
S ns (N body) // th = S ns $ N $ body // th
2023-01-20 20:34:28 -05:00
2023-04-01 13:16:43 -04:00
export %inline CanShift (Term d) where s // by = s // Shift by
export %inline CanShift (Elim d) where e // by = e // Shift by
2023-01-22 21:22:50 -05:00
2023-09-17 07:54:26 -04:00
export %inline CanShift (flip Term n) where s // by = s // Shift by
export %inline CanShift (flip Elim n) where e // by = e // Shift by
2023-01-22 21:22:50 -05:00
export %inline
2023-04-01 13:16:43 -04:00
{s : Nat} -> CanShift (ScopeTermN s d) where
b // by = b // Shift by
2022-04-23 18:21:30 -04:00
export %inline
2023-06-23 12:32:05 -04:00
comp : DSubst d1 d2 -> TSubst d1 n1 mid -> TSubst d2 mid n2 -> TSubst d2 n1 n2
comp th ps ph = map (// th) ps . ph
2022-04-23 18:21:30 -04:00
2023-02-10 15:40:44 -05:00
public export %inline
2023-04-15 09:13:01 -04:00
dweakT : (by : Nat) -> Term d n -> Term (by + d) n
dweakT by t = t // shift by
2023-02-10 15:40:44 -05:00
public export %inline
2023-04-15 09:13:01 -04:00
dweakE : (by : Nat) -> Elim d n -> Elim (by + d) n
dweakE by t = t // shift by
2023-02-10 15:40:44 -05:00
public export %inline
2023-04-15 09:13:01 -04:00
weakT : (by : Nat) -> Term d n -> Term d (by + n)
weakT by t = t // shift by
2023-02-10 15:40:44 -05:00
public export %inline
2023-04-15 09:13:01 -04:00
weakE : (by : Nat) -> Elim d n -> Elim d (by + n)
weakE by t = t // shift by
2023-02-10 15:40:44 -05:00
2023-09-17 07:54:26 -04:00
parameters {auto _ : CanShift f} {s : Nat}
export %inline
getTerm : ScopedBody s f n -> f (s + n)
getTerm (Y b) = b
getTerm (N b) = b // fromNat s
2023-02-22 01:40:19 -05:00
2023-09-17 07:54:26 -04:00
export %inline
(.term) : Scoped s f n -> f (s + n)
t.term = getTerm t.body
2023-02-22 01:40:19 -05:00
2023-09-17 07:54:26 -04:00
namespace ScopeTermBody
export %inline
getTerm0 : ScopedBody 0 f n -> f n
getTerm0 (Y b) = b
getTerm0 (N b) = b
2023-01-20 20:34:28 -05:00
2023-09-17 07:54:26 -04:00
namespace ScopeTermN
export %inline
(.term0) : Scoped 0 f n -> f n
t.term0 = getTerm0 t.body
2023-01-26 13:54:46 -05:00
export %inline
2023-04-01 13:16:43 -04:00
subN : ScopeTermN s d n -> SnocVect s (Elim d n) -> Term d n
2023-03-26 10:09:47 -04:00
subN (S _ (Y body)) es = body // fromSnocVect es
2023-02-22 01:40:19 -05:00
subN (S _ (N body)) _ = body
2023-01-26 13:54:46 -05:00
2023-01-20 20:34:28 -05:00
export %inline
2023-04-01 13:16:43 -04:00
sub1 : ScopeTerm d n -> Elim d n -> Term d n
2023-03-26 10:09:47 -04:00
sub1 t e = subN t [< e]
2023-01-26 13:54:46 -05:00
export %inline
2023-04-01 13:16:43 -04:00
dsubN : DScopeTermN s d n -> SnocVect s (Dim d) -> Term d n
2023-03-26 10:09:47 -04:00
dsubN (S _ (Y body)) ps = body // fromSnocVect ps
2023-02-22 01:40:19 -05:00
dsubN (S _ (N body)) _ = body
2023-01-20 20:34:28 -05:00
export %inline
2023-04-01 13:16:43 -04:00
dsub1 : DScopeTerm d n -> Dim d -> Term d n
2023-03-26 10:09:47 -04:00
dsub1 t p = dsubN t [< p]
2023-01-26 13:54:46 -05:00
2023-01-20 20:34:28 -05:00
public export %inline
2023-11-27 15:01:36 -05:00
(.zero) : (body : DScopeTerm d n) -> {default body.loc loc : Loc} -> Term d n
2023-05-01 21:06:25 -04:00
body.zero = dsub1 body $ K Zero loc
2023-01-20 20:34:28 -05:00
public export %inline
2023-11-27 15:01:36 -05:00
(.one) : (body : DScopeTerm d n) -> {default body.loc loc : Loc} -> Term d n
2023-05-01 21:06:25 -04:00
body.one = dsub1 body $ K One loc
2023-02-26 05:17:42 -05:00
public export
0 CloTest : TermLike -> Type
2023-04-01 13:16:43 -04:00
CloTest tm = forall d, n. tm d n -> Bool
2023-02-26 05:17:42 -05:00
2023-09-20 15:58:55 -04:00
public export
2023-02-26 05:17:42 -05:00
interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where
2023-06-23 12:32:05 -04:00
pushSubstsWith : DSubst d1 d2 -> TSubst d2 n1 n2 ->
tm d1 n1 -> Subset (tm d2 n2) (No . isClo)
2023-02-26 05:17:42 -05:00
public export
2023-04-01 13:16:43 -04:00
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm d n)
2023-02-26 05:17:42 -05:00
NotClo = No . isClo
public export
0 NonClo : (tm : TermLike) -> {isClo : CloTest tm} ->
PushSubsts tm isClo => TermLike
2023-04-01 13:16:43 -04:00
NonClo tm d n = Subset (tm d n) NotClo
2023-02-26 05:17:42 -05:00
public export %inline
nclo : {isClo : CloTest tm} -> (0 _ : PushSubsts tm isClo) =>
2023-04-01 13:16:43 -04:00
(t : tm d n) -> (0 nc : NotClo t) => NonClo tm d n
2023-02-26 05:17:42 -05:00
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
2023-04-01 13:16:43 -04:00
pushSubsts : tm d n -> NonClo tm d n
2023-02-26 05:17:42 -05:00
pushSubsts s = pushSubstsWith id id s
export %inline
2023-06-23 12:32:05 -04:00
pushSubstsWith' : DSubst d1 d2 -> TSubst d2 n1 n2 -> tm d1 n1 -> tm d2 n2
2023-02-26 05:17:42 -05:00
pushSubstsWith' th ph x = fst $ pushSubstsWith th ph x
2023-05-14 13:58:46 -04:00
export %inline
pushSubsts' : tm d n -> tm d n
pushSubsts' s = fst $ pushSubsts s
2023-02-26 05:17:42 -05:00
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
2023-12-04 16:47:52 -05:00
export
PushSubsts Elim Subst.isCloE where
pushSubstsWith th ph (F x u loc) =
nclo $ F x u loc
pushSubstsWith th ph (B i loc) =
let res = getLoc ph i loc in
case nchoose $ isCloE res of
Left yes => assert_total pushSubsts res
Right no => Element res no
pushSubstsWith th ph (App f s loc) =
nclo $ App (f // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CasePair pi p r b loc) =
nclo $ CasePair pi (p // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Fst pair loc) =
nclo $ Fst (pair // th // ph) loc
pushSubstsWith th ph (Snd pair loc) =
nclo $ Snd (pair // th // ph) loc
pushSubstsWith th ph (CaseEnum pi t r arms loc) =
nclo $ CaseEnum pi (t // th // ph) (r // th // ph)
(map (\b => b // th // ph) arms) loc
pushSubstsWith th ph (CaseNat pi pi' n r z s loc) =
nclo $ CaseNat pi pi' (n // th // ph) (r // th // ph)
(z // th // ph) (s // th // ph) loc
pushSubstsWith th ph (CaseBox pi x r b loc) =
nclo $ CaseBox pi (x // th // ph) (r // th // ph) (b // th // ph) loc
pushSubstsWith th ph (DApp f d loc) =
nclo $ DApp (f // th // ph) (d // th) loc
pushSubstsWith th ph (Ann s a loc) =
nclo $ Ann (s // th // ph) (a // th // ph) loc
pushSubstsWith th ph (Coe ty p q val loc) =
nclo $ Coe (ty // th // ph) (p // th) (q // th) (val // th // ph) loc
pushSubstsWith th ph (Comp ty p q val r zero one loc) =
nclo $ Comp (ty // th // ph) (p // th) (q // th)
(val // th // ph) (r // th)
(zero // th // ph) (one // th // ph) loc
pushSubstsWith th ph (TypeCase ty ret arms def loc) =
nclo $ TypeCase (ty // th // ph) (ret // th // ph)
(map (\t => t // th // ph) arms) (def // th // ph) loc
pushSubstsWith th ph (CloE (Sub e ps)) =
pushSubstsWith th (comp th ps ph) e
pushSubstsWith th ph (DCloE (Sub e ps)) =
pushSubstsWith (ps . th) ph e
export
PushSubsts Term Subst.isCloT where
pushSubstsWith th ph (TYPE l loc) =
nclo $ TYPE l loc
pushSubstsWith th ph (IOState loc) =
nclo $ IOState loc
pushSubstsWith th ph (Pi qty a body loc) =
nclo $ Pi qty (a // th // ph) (body // th // ph) loc
pushSubstsWith th ph (Lam body loc) =
nclo $ Lam (body // th // ph) loc
pushSubstsWith th ph (Sig a b loc) =
nclo $ Sig (a // th // ph) (b // th // ph) loc
pushSubstsWith th ph (Pair s t loc) =
nclo $ Pair (s // th // ph) (t // th // ph) loc
pushSubstsWith th ph (Enum tags loc) =
nclo $ Enum tags loc
pushSubstsWith th ph (Tag tag loc) =
nclo $ Tag tag loc
pushSubstsWith th ph (Eq ty l r loc) =
nclo $ Eq (ty // th // ph) (l // th // ph) (r // th // ph) loc
pushSubstsWith th ph (DLam body loc) =
nclo $ DLam (body // th // ph) loc
pushSubstsWith _ _ (NAT loc) =
nclo $ NAT loc
pushSubstsWith _ _ (Nat n loc) =
nclo $ Nat n loc
pushSubstsWith th ph (Succ n loc) =
nclo $ Succ (n // th // ph) loc
pushSubstsWith _ _ (STRING loc) =
nclo $ STRING loc
pushSubstsWith _ _ (Str s loc) =
nclo $ Str s loc
pushSubstsWith th ph (BOX pi ty loc) =
nclo $ BOX pi (ty // th // ph) loc
pushSubstsWith th ph (Box val loc) =
nclo $ Box (val // th // ph) loc
pushSubstsWith th ph (E e) =
let Element e nc = pushSubstsWith th ph e in nclo $ E e
pushSubstsWith th ph (Let qty rhs body loc) =
nclo $ Let qty (rhs // th // ph) (body // th // ph) loc
pushSubstsWith th ph (CloT (Sub s ps)) =
pushSubstsWith th (comp th ps ph) s
pushSubstsWith th ph (DCloT (Sub s ps)) =
pushSubstsWith (ps . th) ph s