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

376 lines
12 KiB
Idris

module Quox.Syntax.Term.Subst
import Quox.No
import Quox.Syntax.Term.Base
import Data.Vect
%default total
namespace CanDSubst
public export
interface CanDSubst (0 tm : Nat -> Nat -> Type) where
(//) : tm dfrom n -> Lazy (DSubst dfrom dto) -> tm dto 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 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
||| 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
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
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
export %inline FromVar (Term q d) where fromVar = E . fromVar
||| 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
export
CanSubstSelf (Elim q d) where
F x // _ = F x
B i // th = th !! i
CloE e ph // th = assert_total CloE e $ ph . th
e // th = case force th of
Shift SZ => e
th => CloE e th
namespace CanTSubst
public export
interface CanTSubst (0 tm : TermLike) where
(//) : tm q d from -> Lazy (TSubst q d from to) -> tm q d to
||| 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 // _ = TYPE l
E e // th = E $ e // th
CloT s ph // th = CloT s $ ph . th
s // th = case force th of
Shift SZ => s
th => CloT s 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
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
export %inline
{s : Nat} -> CanShift (ScopeTermN s q d) where
b // by = b // Shift by
export %inline
comp : DSubst dfrom dto -> TSubst q dfrom from mid -> TSubst q dto mid to ->
TSubst q dto from to
comp th ps ph = map (// th) ps . ph
public export %inline
dweakT : {by : Nat} -> Term q d n -> Term q (by + d) n
dweakT t = t // shift by
public export %inline
dweakE : {by : Nat} -> Elim q d n -> Elim q (by + d) n
dweakE t = t // shift by
public export %inline
weakT : {default 1 by : Nat} -> Term q d n -> Term q d (by + n)
weakT t = t // shift by
public export %inline
weakE : {default 1 by : Nat} -> Elim q d n -> Elim q d (by + n)
weakE t = t // shift by
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 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 (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
sub1 t e = subN t [e]
export %inline
dsubN : DScopeTermN s q d n -> Vect s (Dim d) -> Term q d n
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
dsub1 t p = dsubN t [p]
public export %inline
(.zero) : DScopeTerm q d n -> Term q d n
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
parameters {0 d', n' : Nat}
mutual
namespace Term
export
inject : Term q d n -> Term q (d + d') (n + n')
inject (TYPE l) = TYPE l
inject (Pi qty arg res) = Pi qty (inject arg) (inject res)
inject (Lam body) = Lam (inject body)
inject (Sig fst snd) = Sig (inject fst) (inject snd)
inject (Pair fst snd) = Pair (inject fst) (inject snd)
inject (Enum cases) = Enum cases
inject (Tag tag) = Tag tag
inject (Eq ty l r) = Eq (inject ty) (inject l) (inject r)
inject (DLam body) = DLam (inject body)
inject (E e) = E (inject e)
inject (CloT tm th) = CloT (inject tm) (inject th)
inject (DCloT tm th) = DCloT (inject tm) (inject th)
namespace Elim
export
inject : Elim q d n -> Elim q (d + d') (n + n')
inject (F x) = F x
inject (B i) = B $ inject i
inject (fun :@ arg) = (inject fun) :@ (inject arg)
inject (CasePair qty pair ret body) =
CasePair qty (inject pair) (inject ret) (inject body)
inject (CaseEnum qty tag ret arms) =
CaseEnum qty (inject tag) (inject ret)
(assert_total $ map inject arms)
inject (fun :% arg) = (inject fun) :% (inject arg)
inject (tm :# ty) = (inject tm) :# (inject ty)
inject (CloE el th) = CloE (inject el) (inject th)
inject (DCloE el th) = DCloE (inject el) (inject th)
namespace ScopeTerm
export
inject : ScopeTermN s q d n -> ScopeTermN s q (d + d') (n + n')
inject (S names (N body)) = S names (N (inject body))
inject (S names (Y body)) {s, n} =
S names $ Y $ rewrite plusAssociative s n n' in inject body
namespace DScopeTerm
export
inject : DScopeTermN s q d n -> DScopeTermN s q (d + d') (n + n')
inject (S names (N body)) = S names (N (inject body))
inject (S names (Y body)) {s, d} =
S names $ Y $ rewrite plusAssociative s d d' in inject body
namespace TSubst
export
inject : TSubst q d from to -> TSubst q (d + d') (from + n') (to + n')
inject (Shift by) = Shift $ inject by
inject (t ::: th) = inject t ::: inject th
namespace DSubst
export
inject : DSubst from to -> DSubst (from + d') (to + d')
inject (Shift by) = Shift $ inject by
inject (p ::: th) = inject p ::: inject th
-- [fixme]
-- Error: Linearity checking failed on metavar Quox.Syntax.Term.Subst.{b:7362}
-- (Int not a function type)
{-
%transform "Term.inject" Term.inject t = believe_me t
%transform "Elim.inject" Elim.inject e = believe_me e
%transform "ScopeTerm.inject" ScopeTerm.inject t = believe_me t
%transform "DScopeTerm.inject" DScopeTerm.inject t = believe_me t
%transform "TSubst.inject" TSubst.inject th = believe_me th
%transform "DSubst.inject" DSubst.inject th = believe_me th
-}