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
|
|
|
|
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
|
|
|
|
2023-02-20 16:22:23 -05:00
|
|
|
namespace CanDSubst
|
|
|
|
public export
|
2023-04-01 13:16:43 -04:00
|
|
|
interface CanDSubst (0 tm : TermLike) where
|
2024-05-27 15:28:22 -04:00
|
|
|
(//) : tm q d1 n -> Lazy (DSubst d1 d2) -> tm q 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
|
2023-04-27 15:37:20 -04:00
|
|
|
s // Shift SZ = s
|
2023-05-01 21:06:25 -04:00
|
|
|
TYPE l loc // _ = TYPE l loc
|
2023-04-27 15:37:20 -04:00
|
|
|
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
|
2024-05-27 15:28:22 -04:00
|
|
|
subDArgs : Elim q d1 n -> DSubst d1 d2 -> Elim q 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
|
2023-04-27 15:37:20 -04:00
|
|
|
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
|
2023-04-27 15:37:20 -04:00
|
|
|
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
|
|
|
|
e // th = DCloE $ Sub e th
|
2023-02-22 01:40:19 -05:00
|
|
|
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
|
|
|
|
2023-02-22 01:40:19 -05:00
|
|
|
namespace DSubst.ScopeTermN
|
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
(//) : 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
|
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} ->
|
2024-05-27 15:28:22 -04:00
|
|
|
DScopeTermN s q d1 n -> Lazy (DSubst d1 d2) ->
|
|
|
|
DScopeTermN s q 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
|
|
|
|
|
|
|
|
2024-05-23 18:23:09 -04:00
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
FromVarR (Elim q d) where fromVarR = B
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2024-05-23 18:23:09 -04:00
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
FromVar (Elim q d) where fromVar = B; fromVarSame _ _ = Refl
|
2024-05-23 18:23:09 -04:00
|
|
|
|
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
FromVarR (Term q d) where fromVarR = E .: fromVarR
|
2024-05-23 18:23:09 -04:00
|
|
|
|
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
FromVar (Term q d) where fromVar = E .: fromVar; fromVarSame _ _ = Refl
|
2024-05-23 18:23:09 -04:00
|
|
|
|
|
|
|
|
|
|
|
export
|
2024-05-27 15:28:22 -04:00
|
|
|
CanSubstSelf (Elim q d)
|
2024-05-23 18:23:09 -04:00
|
|
|
|
|
|
|
private
|
2024-05-27 15:28:22 -04:00
|
|
|
tsubstElim : Elim q d from -> Lazy (TSubst q d from to) -> Elim q d to
|
2024-05-23 18:23:09 -04:00
|
|
|
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
|
2023-02-20 16:22:23 -05: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
|
2024-05-27 15:28:22 -04:00
|
|
|
CanSubstSelfR (Elim q d) where (//?) = tsubstElim
|
2024-05-23 18:23:09 -04:00
|
|
|
|
2022-04-23 18:21:30 -04:00
|
|
|
export
|
2024-05-27 15:28:22 -04:00
|
|
|
CanSubstSelf (Elim q d) where (//) = tsubstElim; substSame _ _ = Refl
|
2022-04-23 18:21:30 -04:00
|
|
|
|
2023-02-20 16:22:23 -05:00
|
|
|
namespace CanTSubst
|
|
|
|
public export
|
|
|
|
interface CanTSubst (0 tm : TermLike) where
|
2024-05-27 15:28:22 -04:00
|
|
|
(//) : tm q d n1 -> Lazy (TSubst q d n1 n2) -> tm q d n2
|
2023-02-20 16:22:23 -05:00
|
|
|
|
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
|
2023-02-20 16:22:23 -05:00
|
|
|
CanTSubst Term where
|
2023-05-01 21:06:25 -04:00
|
|
|
TYPE l loc // _ = TYPE l loc
|
2023-04-27 15:37:20 -04:00
|
|
|
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} ->
|
2024-05-27 15:28:22 -04:00
|
|
|
ScopeTermN s q d n1 -> Lazy (TSubst q d n1 n2) ->
|
|
|
|
ScopeTermN s q 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} ->
|
2024-05-27 15:28:22 -04:00
|
|
|
DScopeTermN s q d n1 -> Lazy (TSubst q d n1 n2) ->
|
|
|
|
DScopeTermN s q 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
|
|
|
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
2023-01-22 21:22:50 -05:00
|
|
|
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
2023-09-17 07:54:26 -04:00
|
|
|
|
2023-01-22 21:22:50 -05:00
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
{s : Nat} -> CanShift (ScopeTermN s q d) where
|
2023-02-20 16:22:23 -05:00
|
|
|
b // by = b // Shift by
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
2022-04-23 18:21:30 -04:00
|
|
|
|
|
|
|
|
2023-02-10 15:40:44 -05:00
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
dweakT : (by : Nat) -> Term q d n -> Term q (by + d) n
|
2023-04-15 09:13:01 -04:00
|
|
|
dweakT by t = t // shift by
|
2023-02-10 15:40:44 -05:00
|
|
|
|
2024-02-24 10:04:38 -05:00
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
dweakS : (by : Nat) -> ScopeTermN s q d n -> ScopeTermN s q (by + d) n
|
2024-02-24 10:04:38 -05:00
|
|
|
dweakS by t = t // shift by
|
|
|
|
|
|
|
|
public export %inline
|
|
|
|
dweakDS : {s : Nat} -> (by : Nat) ->
|
2024-05-27 15:28:22 -04:00
|
|
|
DScopeTermN s q d n -> DScopeTermN s q (by + d) n
|
2024-02-24 10:04:38 -05:00
|
|
|
dweakDS by t = t // shift by
|
|
|
|
|
2023-02-10 15:40:44 -05:00
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
dweakE : (by : Nat) -> Elim q d n -> Elim q (by + d) n
|
2023-04-15 09:13:01 -04:00
|
|
|
dweakE by t = t // shift by
|
2023-02-10 15:40:44 -05:00
|
|
|
|
|
|
|
|
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
weakT : (by : Nat) -> Term q d n -> Term q d (by + n)
|
2023-04-15 09:13:01 -04:00
|
|
|
weakT by t = t // shift by
|
2023-02-10 15:40:44 -05:00
|
|
|
|
2024-02-24 10:04:38 -05:00
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
weakS : {s : Nat} -> (by : Nat) ->
|
|
|
|
ScopeTermN s q d n -> ScopeTermN s q d (by + n)
|
2024-02-24 10:04:38 -05:00
|
|
|
weakS by t = t // shift by
|
|
|
|
|
|
|
|
public export %inline
|
|
|
|
weakDS : {s : Nat} -> (by : Nat) ->
|
2024-05-27 15:28:22 -04:00
|
|
|
DScopeTermN s q d n -> DScopeTermN s q d (by + n)
|
2024-02-24 10:04:38 -05:00
|
|
|
weakDS by t = t // shift by
|
|
|
|
|
2023-02-10 15:40:44 -05:00
|
|
|
public export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
weakE : (by : Nat) -> Elim q d n -> Elim q d (by + n)
|
2023-04-15 09:13:01 -04:00
|
|
|
weakE by t = t // shift by
|
2023-02-10 15:40:44 -05:00
|
|
|
|
2024-05-27 15:28:22 -04:00
|
|
|
-- no weakQ etc because no first-class binder to push under
|
|
|
|
|
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
|
2024-05-27 15:28:22 -04:00
|
|
|
subN : ScopeTermN s q d n -> SnocVect s (Elim q d n) -> Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
sub1 : ScopeTerm q d n -> Elim q d n -> Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
dsubN : DScopeTermN s q d n -> SnocVect s (Dim d) -> Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
dsub1 : DScopeTerm q d n -> Dim d -> Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
(.zero) : (body : DScopeTerm q d n) -> {default body.loc loc : Loc} ->
|
|
|
|
Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
(.one) : (body : DScopeTerm q d n) -> {default body.loc loc : Loc} ->
|
|
|
|
Term q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
CloTest tm = forall q, d, n. tm q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
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)
|
2023-02-26 05:17:42 -05:00
|
|
|
|
|
|
|
public export
|
2024-05-27 15:28:22 -04:00
|
|
|
0 NotClo : {isClo : CloTest tm} -> PushSubsts tm isClo => Pred (tm q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
NonClo tm q d n = Subset (tm q d n) NotClo
|
2023-02-26 05:17:42 -05:00
|
|
|
|
|
|
|
public export %inline
|
|
|
|
nclo : {isClo : CloTest tm} -> (0 _ : PushSubsts tm isClo) =>
|
2024-05-27 15:28:22 -04:00
|
|
|
(t : tm q d n) -> (0 nc : NotClo t) => NonClo tm q 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
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubsts : {q : Nat} -> tm q d n -> NonClo tm q d n
|
|
|
|
pushSubsts s = pushSubstsWith id id id s
|
2023-02-26 05:17:42 -05:00
|
|
|
|
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
2023-02-26 05:17:42 -05:00
|
|
|
|
2023-05-14 13:58:46 -04:00
|
|
|
export %inline
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubsts' : {q : Nat} -> tm q d n -> tm q d n
|
2023-05-14 13:58:46 -04:00
|
|
|
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
|
2024-05-27 15:28:22 -04:00
|
|
|
isCloT (QCloT {}) = True
|
2023-02-26 05:17:42 -05:00
|
|
|
isCloT (E e) = isCloE e
|
|
|
|
isCloT _ = False
|
|
|
|
|
|
|
|
public export
|
|
|
|
isCloE : CloTest Elim
|
|
|
|
isCloE (CloE {}) = True
|
|
|
|
isCloE (DCloE {}) = True
|
2024-05-27 15:28:22 -04:00
|
|
|
isCloE (QCloE {}) = True
|
2023-02-26 05:17:42 -05:00
|
|
|
isCloE _ = False
|
|
|
|
|
2023-12-04 16:47:52 -05:00
|
|
|
export
|
|
|
|
PushSubsts Elim Subst.isCloE where
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith th ph ps (F x u loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ F x u loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith th ph ps (B i loc) =
|
|
|
|
let res = get ps i loc in
|
2023-12-04 16:47:52 -05:00
|
|
|
case nchoose $ isCloE res of
|
|
|
|
Left yes => assert_total pushSubsts res
|
|
|
|
Right no => Element res no
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|
2023-12-04 16:47:52 -05:00
|
|
|
|
|
|
|
export
|
|
|
|
PushSubsts Term Subst.isCloT where
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith th ph ps (TYPE l loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ TYPE l loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith _ _ _ (IOState loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ IOState loc
|
2024-05-27 15:28:22 -04:00
|
|
|
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) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ Enum tags loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith th ph ps (Tag tag loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ Tag tag loc
|
2024-05-27 15:28:22 -04:00
|
|
|
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) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ NAT loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith _ _ _ (Nat n loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ Nat n loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith th ph ps (Succ n loc) =
|
|
|
|
nclo $ Succ (n // th // ph // ps) loc
|
|
|
|
pushSubstsWith _ _ _ (STRING loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ STRING loc
|
2024-05-27 15:28:22 -04:00
|
|
|
pushSubstsWith _ _ _ (Str s loc) =
|
2023-12-04 16:47:52 -05:00
|
|
|
nclo $ Str s loc
|
2024-05-27 15:28:22 -04:00
|
|
|
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
|