module Quox.Syntax.Term.Subst import Quox.No import Quox.Syntax.Term.Base import Data.SnocVect %default total namespace CanQSubst public export interface CanQSubst (0 tm : TermLike) where (//) : {q1, q2 : Nat} -> tm q1 d n -> Lazy (QSubst q1 q2) -> tm q2 d n ||| 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 CanQSubst Term where s // Shift SZ = s TYPE l loc // _ = TYPE l loc QCloT (SubR s ph) // th = QCloT $ SubR s $ ph .? th s // th = QCloT $ SubR s 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 CanQSubst Elim where e // Shift SZ = e F x u loc // _ = F x u loc -- [todo] q args B i loc // _ = B i loc QCloE (SubR e ph) // th = QCloE $ SubR e $ ph .? th e // th = QCloE $ SubR e th namespace CanDSubst public export interface CanDSubst (0 tm : TermLike) where (//) : tm q d1 n -> Lazy (DSubst d1 d2) -> tm q d2 n ||| 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 where s // Shift SZ = s TYPE l loc // _ = TYPE l loc DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th s // th = DCloT $ Sub s th private subDArgs : Elim q d1 n -> DSubst d1 d2 -> Elim q d2 n subDArgs (DApp f d loc) th = DApp (subDArgs f th) (d // th) loc subDArgs e th = DCloE $ Sub 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 where e // Shift SZ = e F x u loc // _ = F x u loc 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 namespace QSubst.ScopeTermN export %inline (//) : {q1, q2 : Nat} -> ScopeTermN s q1 d n -> Lazy (QSubst q1 q2) -> ScopeTermN s q2 d n S ns (Y body) // th = S ns $ Y $ body // th S ns (N body) // th = S ns $ N $ body // th namespace DSubst.ScopeTermN export %inline (//) : ScopeTermN s q d1 n -> Lazy (DSubst d1 d2) -> ScopeTermN s q d2 n S ns (Y body) // th = S ns $ Y $ body // th S ns (N body) // th = S ns $ N $ body // th namespace QSubst.DScopeTermN export %inline (//) : {q1, q2 : Nat} -> DScopeTermN s q1 d n -> Lazy (QSubst q1 q2) -> DScopeTermN s q2 d n 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} -> DScopeTermN s q d1 n -> Lazy (DSubst d1 d2) -> DScopeTermN s q d2 n S ns (Y body) // th = S ns $ Y $ body // pushN s (locs $ toList' ns) th S ns (N body) // th = S ns $ N $ body // th export %inline FromVarR (Elim q d) where fromVarR = B export %inline FromVar (Elim q d) where fromVar = B; fromVarSame _ _ = Refl export %inline FromVarR (Term q d) where fromVarR = E .: fromVarR export %inline FromVar (Term q d) where fromVar = E .: fromVar; fromVarSame _ _ = Refl export CanSubstSelf (Elim q d) private tsubstElim : Elim q d from -> Lazy (TSubst q d from to) -> Elim q d to tsubstElim (F x u loc) _ = F x u loc tsubstElim (B i loc) th = get th i loc tsubstElim (CloE (Sub e ph)) th = assert_total CloE $ Sub e $ ph . th tsubstElim e th = case force th of Shift SZ => e th => CloE $ Sub e th ||| does the minimal reasonable work: ||| - deletes the closure around a *free* name ||| - deletes an identity substitution ||| - composes (lazily) with an existing top-level closure ||| - immediately looks up a bound variable ||| - otherwise, wraps in a new closure CanSubstSelfR (Elim q d) where (//?) = tsubstElim export CanSubstSelf (Elim q d) where (//) = tsubstElim; substSame _ _ = Refl namespace CanTSubst public export interface CanTSubst (0 tm : TermLike) where (//) : tm q d n1 -> Lazy (TSubst q d n1 n2) -> tm q d n2 ||| 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 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 namespace ScopeTermN export %inline (//) : {s : Nat} -> ScopeTermN s q d n1 -> Lazy (TSubst q d n1 n2) -> ScopeTermN s q d n2 S ns (Y body) // th = S ns $ Y $ body // pushN s (locs $ toList' ns) th S ns (N body) // th = S ns $ N $ body // th namespace DScopeTermN export %inline (//) : {s : Nat} -> DScopeTermN s q d n1 -> Lazy (TSubst q d n1 n2) -> DScopeTermN s q d n2 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 -- -- from is not accessible in this context -- export %inline CanShift (\q => Term q d n) where s // by = s // Shift by -- export %inline CanShift (\q => Elim q d n) where e // by = e // Shift by export %inline CanShift (\d => Term q d n) where s // by = s // Shift by export %inline CanShift (\d => Elim q d n) where e // by = e // Shift by export %inline {s : Nat} -> CanShift (ScopeTermN s q d) where b // by = b // Shift by export %inline comp : {q1, q2 : Nat} -> QSubst q1 q2 -> DSubst d1 d2 -> TSubst q1 d1 n1 mid -> TSubst q2 d2 mid n2 -> TSubst q2 d2 n1 n2 comp th ph ps ps' = map (\t => t // th // ph) ps . ps' -- export %inline -- compD : DSubst d1 d2 -> TSubst q d1 n1 mid -> -- TSubst q d2 mid n2 -> TSubst q d2 n1 n2 -- compD th ps ph = map (// th) ps . ph -- export %inline -- compQ : {q1, q2 : Nat} -> QSubst q1 q2 -> TSubst q1 d n1 mid -> -- TSubst q2 d mid n2 -> TSubst q2 d n1 n2 -- compQ th ps ph = map (// th) ps . ph public export %inline dweakT : (by : Nat) -> Term q d n -> Term q (by + d) n dweakT by t = t // shift by public export %inline dweakS : (by : Nat) -> ScopeTermN s q d n -> ScopeTermN s q (by + d) n dweakS by t = t // shift by public export %inline dweakDS : {s : Nat} -> (by : Nat) -> DScopeTermN s q d n -> DScopeTermN s q (by + d) n dweakDS by t = t // shift by public export %inline dweakE : (by : Nat) -> Elim q d n -> Elim q (by + d) n dweakE by t = t // shift by public export %inline weakT : (by : Nat) -> Term q d n -> Term q d (by + n) weakT by t = t // shift by public export %inline weakS : {s : Nat} -> (by : Nat) -> ScopeTermN s q d n -> ScopeTermN s q d (by + n) weakS by t = t // shift by public export %inline weakDS : {s : Nat} -> (by : Nat) -> DScopeTermN s q d n -> DScopeTermN s q d (by + n) weakDS by t = t // shift by public export %inline weakE : (by : Nat) -> Elim q d n -> Elim q d (by + n) weakE by t = t // shift by -- no weakQ etc because no first-class binder to push under 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 export %inline (.term) : Scoped s f n -> f (s + n) t.term = getTerm t.body namespace ScopeTermBody export %inline getTerm0 : ScopedBody 0 f n -> f n getTerm0 (Y b) = b getTerm0 (N b) = b namespace ScopeTermN export %inline (.term0) : Scoped 0 f n -> f n t.term0 = getTerm0 t.body export %inline subN : ScopeTermN s q d n -> SnocVect s (Elim q d n) -> Term q d n subN (S _ (Y body)) es = body // fromSnocVect es subN (S _ (N body)) _ = body export %inline sub1 : ScopeTerm q d n -> Elim q d n -> Term q d n sub1 t e = subN t [< e] export %inline dsubN : DScopeTermN s q d n -> SnocVect s (Dim d) -> Term q d n dsubN (S _ (Y body)) ps = body // fromSnocVect ps dsubN (S _ (N body)) _ = body export %inline dsub1 : DScopeTerm q d n -> Dim d -> Term q d n dsub1 t p = dsubN t [< p] public export %inline (.zero) : (body : DScopeTerm q d n) -> {default body.loc loc : Loc} -> Term q d n body.zero = dsub1 body $ K Zero loc public export %inline (.one) : (body : DScopeTerm q d n) -> {default body.loc loc : Loc} -> Term q d n body.one = dsub1 body $ K One loc public export 0 CloTest : TermLike -> Type CloTest tm = forall q, d, n. tm q d n -> Bool public export interface PushSubsts (0 tm : TermLike) (0 isClo : CloTest tm) | tm where pushSubstsWith : {q1, q2 : Nat} -> QSubst q1 q2 -> DSubst d1 d2 -> TSubst q2 d2 n1 n2 -> tm q1 d1 n1 -> Subset (tm q2 d2 n2) (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 : {q : Nat} -> tm q d n -> NonClo tm q d n pushSubsts s = pushSubstsWith id id id s export %inline pushSubstsWith' : {q1, q2 : Nat} -> QSubst q1 q2 -> DSubst d1 d2 -> TSubst q2 d2 n1 n2 -> tm q1 d1 n1 -> tm q2 d2 n2 pushSubstsWith' th ph ps x = fst $ pushSubstsWith th ph ps x export %inline pushSubsts' : {q : Nat} -> tm q d n -> tm q d n pushSubsts' s = fst $ pushSubsts s mutual public export isCloT : CloTest Term isCloT (CloT {}) = True isCloT (DCloT {}) = True isCloT (QCloT {}) = True isCloT (E e) = isCloE e isCloT _ = False public export isCloE : CloTest Elim isCloE (CloE {}) = True isCloE (DCloE {}) = True isCloE (QCloE {}) = True isCloE _ = False export PushSubsts Elim Subst.isCloE where pushSubstsWith th ph ps (F x u loc) = nclo $ F x u loc pushSubstsWith th ph ps (B i loc) = let res = get ps i loc in case nchoose $ isCloE res of Left yes => assert_total pushSubsts res Right no => Element res no pushSubstsWith th ph ps (App f s loc) = nclo $ App (f // th // ph // ps) (s // th // ph // ps) loc pushSubstsWith th ph ps (CasePair pi p r b loc) = nclo $ CasePair (pi //? th) (p // th // ph // ps) (r // th // ph // ps) (b // th // ph // ps) loc pushSubstsWith th ph ps (Fst pair loc) = nclo $ Fst (pair // th // ph // ps) loc pushSubstsWith th ph ps (Snd pair loc) = nclo $ Snd (pair // th // ph // ps) loc pushSubstsWith th ph ps (CaseEnum pi t r arms loc) = nclo $ CaseEnum (pi //? th) (t // th // ph // ps) (r // th // ph // ps) (map (\b => b // th // ph // ps) arms) loc pushSubstsWith th ph ps (CaseNat pi pi' n r z s loc) = nclo $ CaseNat (pi //? th) (pi' //? th) (n // th // ph // ps) (r // th // ph // ps) (z // th // ph // ps) (s // th // ph // ps) loc pushSubstsWith th ph ps (CaseBox pi x r b loc) = nclo $ CaseBox (pi //? th) (x // th // ph // ps) (r // th // ph // ps) (b // th // ph // ps) loc pushSubstsWith th ph ps (DApp f d loc) = nclo $ DApp (f // th // ph // ps) (d // ph) loc pushSubstsWith th ph ps (Ann s a loc) = nclo $ Ann (s // th // ph // ps) (a // th // ph // ps) loc pushSubstsWith th ph ps (Coe ty p q val loc) = nclo $ Coe (ty // th // ph // ps) (p // ph) (q // ph) (val // th // ph // ps) loc pushSubstsWith th ph ps (Comp ty p q val r zero one loc) = nclo $ Comp (ty // th // ph // ps) (p // ph) (q // ph) (val // th // ph // ps) (r // ph) (zero // th // ph // ps) (one // th // ph // ps) loc pushSubstsWith th ph ps (TypeCase ty ret arms def loc) = nclo $ TypeCase (ty // th // ph // ps) (ret // th // ph // ps) (map (\t => t // th // ph // ps) arms) (def // th // ph // ps) loc pushSubstsWith th ph ps (CloE (Sub e ps')) = pushSubstsWith th ph (comp th ph ps' ps) e pushSubstsWith th ph ps (DCloE (Sub e ph')) = pushSubstsWith th (ph' . ph) ps e pushSubstsWith th ph ps (QCloE (SubR e th')) = pushSubstsWith (th' .? th) ph ps e export PushSubsts Term Subst.isCloT where pushSubstsWith th ph ps (TYPE l loc) = nclo $ TYPE l loc pushSubstsWith _ _ _ (IOState loc) = nclo $ IOState loc pushSubstsWith th ph ps (Pi qty a body loc) = nclo $ Pi (qty //? th) (a // th // ph // ps) (body // th // ph // ps) loc pushSubstsWith th ph ps (Lam body loc) = nclo $ Lam (body // th // ph // ps) loc pushSubstsWith th ph ps (Sig a b loc) = nclo $ Sig (a // th // ph // ps) (b // th // ph // ps) loc pushSubstsWith th ph ps (Pair s t loc) = nclo $ Pair (s // th // ph // ps) (t // th // ph // ps) loc pushSubstsWith th ph ps (Enum tags loc) = nclo $ Enum tags loc pushSubstsWith th ph ps (Tag tag loc) = nclo $ Tag tag loc pushSubstsWith th ph ps (Eq ty l r loc) = nclo $ Eq (ty // th // ph // ps) (l // th // ph // ps) (r // th // ph // ps) loc pushSubstsWith th ph ps (DLam body loc) = nclo $ DLam (body // th // ph // ps) loc pushSubstsWith _ _ _ (NAT loc) = nclo $ NAT loc pushSubstsWith _ _ _ (Nat n loc) = nclo $ Nat n loc pushSubstsWith th ph ps (Succ n loc) = nclo $ Succ (n // th // ph // ps) loc pushSubstsWith _ _ _ (STRING loc) = nclo $ STRING loc pushSubstsWith _ _ _ (Str s loc) = nclo $ Str s loc pushSubstsWith th ph ps (BOX pi ty loc) = nclo $ BOX (pi //? th) (ty // th // ph // ps) loc pushSubstsWith th ph ps (Box val loc) = nclo $ Box (val // th // ph // ps) loc pushSubstsWith th ph ps (E e) = let Element e nc = pushSubstsWith th ph ps e in nclo $ E e pushSubstsWith th ph ps (Let qty rhs body loc) = nclo $ Let (qty //? th) (rhs // th // ph // ps) (body // th // ph // ps) loc pushSubstsWith th ph ps (CloT (Sub s ps')) = pushSubstsWith th ph (comp th ph ps' ps) s pushSubstsWith th ph ps (DCloT (Sub s ph')) = pushSubstsWith th (ph' . ph) ps s pushSubstsWith th ph ps (QCloT (SubR s th')) = pushSubstsWith (th' .? th) ph ps s