rhiannon morris
3fb8580f85
e.g. "coe [_ ⇒ A] @p @q s" should immediately reduce to "s", but if the "_ ⇒ A" happened to use an SY it didn't. this will still happen if a wrong SY sneaks in but the alternative is re-traversing the term over and over every time whnf runs
274 lines
9.5 KiB
Idris
274 lines
9.5 KiB
Idris
module Quox.Syntax.Term.Tighten
|
|
|
|
import Quox.Syntax.Term.Base
|
|
import Quox.Syntax.Subst
|
|
import public Quox.OPE
|
|
|
|
%default total
|
|
|
|
|
|
export
|
|
Tighten (Shift from) where
|
|
-- `OPE m n` is a spicy `m ≤ n`,
|
|
-- and `Shift from n` is a (different) spicy `from ≤ n`
|
|
-- so the value is `from ≤ m` (as a `Shift`), if that is the case
|
|
tighten _ SZ = Nothing
|
|
tighten Id by = Just by
|
|
tighten (Drop p) (SS by) = tighten p by
|
|
tighten (Keep p) (SS by) = [|SS $ tighten p by|]
|
|
|
|
|
|
export
|
|
tightenSub : (forall m, n. OPE m n -> env n -> Maybe (env m)) ->
|
|
OPE to1 to2 -> Subst env from to2 -> Maybe (Subst env from to1)
|
|
tightenSub f p (Shift by) = [|Shift $ tighten p by|]
|
|
tightenSub f p (t ::: th) = [|f p t !::: tightenSub f p th|]
|
|
|
|
export
|
|
Tighten env => Tighten (Subst env from) where
|
|
tighten p th = tightenSub tighten p th
|
|
|
|
|
|
export
|
|
tightenScope : (forall m, n. OPE m n -> f n -> Maybe (f m)) ->
|
|
{s : Nat} -> OPE m n -> Scoped s f n -> Maybe (Scoped s f m)
|
|
tightenScope f p (S names (Y body)) = SY names <$> f (keepN s p) body
|
|
tightenScope f p (S names (N body)) = S names . N <$> f p body
|
|
|
|
export
|
|
tightenDScope : {0 f : Nat -> Nat -> Type} ->
|
|
(forall m, n, k. OPE m n -> f n k -> Maybe (f m k)) ->
|
|
OPE m n -> Scoped s (f n) k -> Maybe (Scoped s (f m) k)
|
|
tightenDScope f p (S names (Y body)) = SY names <$> f p body
|
|
tightenDScope f p (S names (N body)) = S names . N <$> f p body
|
|
|
|
|
|
mutual
|
|
private
|
|
tightenT : OPE n1 n2 -> Term d n2 -> Maybe (Term d n1)
|
|
tightenT p (TYPE l) = pure $ TYPE l
|
|
tightenT p (Pi qty arg res) =
|
|
Pi qty <$> tightenT p arg <*> tightenS p res
|
|
tightenT p (Lam body) = Lam <$> tightenS p body
|
|
tightenT p (Sig fst snd) = Sig <$> tightenT p fst <*> tightenS p snd
|
|
tightenT p (Pair fst snd) = Pair <$> tightenT p fst <*> tightenT p snd
|
|
tightenT p (Enum cases) = pure $ Enum cases
|
|
tightenT p (Tag tag) = pure $ Tag tag
|
|
tightenT p (Eq ty l r) =
|
|
Eq <$> tightenDS p ty <*> tightenT p l <*> tightenT p r
|
|
tightenT p (DLam body) = DLam <$> tightenDS p body
|
|
tightenT p Nat = pure Nat
|
|
tightenT p Zero = pure Zero
|
|
tightenT p (Succ s) = Succ <$> tightenT p s
|
|
tightenT p (BOX qty ty) = BOX qty <$> tightenT p ty
|
|
tightenT p (Box val) = Box <$> tightenT p val
|
|
tightenT p (E e) = assert_total $ E <$> tightenE p e
|
|
tightenT p (CloT tm th) = do
|
|
th <- assert_total $ tightenSub tightenE p th
|
|
pure $ CloT tm th
|
|
tightenT p (DCloT tm th) = [|DCloT (tightenT p tm) (pure th)|]
|
|
|
|
private
|
|
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
|
|
tightenE p (F x) = pure $ F x
|
|
tightenE p (B i) = [|B $ tighten p i|]
|
|
tightenE p (fun :@ arg) = [|tightenE p fun :@ tightenT p arg|]
|
|
tightenE p (CasePair qty pair ret body) =
|
|
CasePair qty <$> tightenE p pair
|
|
<*> tightenS p ret
|
|
<*> tightenS p body
|
|
tightenE p (CaseEnum qty tag ret arms) =
|
|
CaseEnum qty <$> tightenE p tag
|
|
<*> tightenS p ret
|
|
<*> traverse (tightenT p) arms
|
|
tightenE p (CaseNat qty qtyIH nat ret zero succ) =
|
|
CaseNat qty qtyIH
|
|
<$> tightenE p nat
|
|
<*> tightenS p ret
|
|
<*> tightenT p zero
|
|
<*> tightenS p succ
|
|
tightenE p (CaseBox qty box ret body) =
|
|
CaseBox qty <$> tightenE p box
|
|
<*> tightenS p ret
|
|
<*> tightenS p body
|
|
tightenE p (fun :% arg) = (:% arg) <$> tightenE p fun
|
|
tightenE p (tm :# ty) = [|tightenT p tm :# tightenT p ty|]
|
|
tightenE p (Coe ty q0 q1 val) =
|
|
Coe <$> tightenDS p ty
|
|
<*> pure q0 <*> pure q1
|
|
<*> tightenT p val
|
|
tightenE p (Comp ty q0 q1 val r zero one) =
|
|
Comp <$> tightenT p ty
|
|
<*> pure q0 <*> pure q1
|
|
<*> tightenT p val
|
|
<*> pure r
|
|
<*> tightenDS p zero
|
|
<*> tightenDS p one
|
|
tightenE p (TypeCase ty ret arms def) =
|
|
TypeCase <$> tightenE p ty
|
|
<*> tightenT p ret
|
|
<*> traverse (tightenS p) arms
|
|
<*> tightenT p def
|
|
tightenE p (CloE el th) = do
|
|
th <- assert_total $ tightenSub tightenE p th
|
|
pure $ CloE el th
|
|
tightenE p (DCloE el th) = [|DCloE (tightenE p el) (pure th)|]
|
|
|
|
export
|
|
tightenS : {s : Nat} -> OPE m n ->
|
|
ScopeTermN s f n -> Maybe (ScopeTermN s f m)
|
|
tightenS = assert_total $ tightenScope tightenT
|
|
|
|
export
|
|
tightenDS : OPE m n -> DScopeTermN s f n -> Maybe (DScopeTermN s f m)
|
|
tightenDS = assert_total $ tightenDScope tightenT {f = \n, d => Term d n}
|
|
|
|
export Tighten (Elim d) where tighten p e = tightenE p e
|
|
export Tighten (Term d) where tighten p t = tightenT p t
|
|
|
|
export
|
|
Tighten Dim where
|
|
tighten p (K e) = pure $ K e
|
|
tighten p (B i) = B <$> tighten p i
|
|
|
|
|
|
mutual
|
|
export
|
|
dtightenT : OPE d1 d2 -> Term d2 n -> Maybe (Term d1 n)
|
|
dtightenT p (TYPE l) = pure $ TYPE l
|
|
dtightenT p (Pi qty arg res) =
|
|
Pi qty <$> dtightenT p arg <*> dtightenS p res
|
|
dtightenT p (Lam body) =
|
|
Lam <$> dtightenS p body
|
|
dtightenT p (Sig fst snd) =
|
|
Sig <$> dtightenT p fst <*> dtightenS p snd
|
|
dtightenT p (Pair fst snd) =
|
|
Pair <$> dtightenT p fst <*> dtightenT p snd
|
|
dtightenT p (Enum cases) = pure $ Enum cases
|
|
dtightenT p (Tag tag) = pure $ Tag tag
|
|
dtightenT p (Eq ty l r) =
|
|
Eq <$> dtightenDS p ty <*> dtightenT p l <*> dtightenT p r
|
|
dtightenT p (DLam body) = DLam <$> dtightenDS p body
|
|
dtightenT p Nat = pure Nat
|
|
dtightenT p Zero = pure Zero
|
|
dtightenT p (Succ s) = Succ <$> dtightenT p s
|
|
dtightenT p (BOX qty ty) = BOX qty <$> dtightenT p ty
|
|
dtightenT p (Box val) = Box <$> dtightenT p val
|
|
dtightenT p (E e) = assert_total $ E <$> dtightenE p e
|
|
dtightenT p (CloT tm th) = do
|
|
tm <- dtightenT p tm
|
|
th <- assert_total $ traverse (dtightenE p) th
|
|
pure $ CloT tm th
|
|
dtightenT p (DCloT tm th) = do th <- tighten p th; pure $ DCloT tm th
|
|
|
|
export
|
|
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
|
|
dtightenE p (F x) = pure $ F x
|
|
dtightenE p (B i) = pure $ B i
|
|
dtightenE p (fun :@ arg) = [|dtightenE p fun :@ dtightenT p arg|]
|
|
dtightenE p (CasePair qty pair ret body) =
|
|
CasePair qty <$> dtightenE p pair
|
|
<*> dtightenS p ret
|
|
<*> dtightenS p body
|
|
dtightenE p (CaseEnum qty tag ret arms) =
|
|
CaseEnum qty <$> dtightenE p tag
|
|
<*> dtightenS p ret
|
|
<*> traverse (dtightenT p) arms
|
|
dtightenE p (CaseNat qty qtyIH nat ret zero succ) =
|
|
CaseNat qty qtyIH
|
|
<$> dtightenE p nat
|
|
<*> dtightenS p ret
|
|
<*> dtightenT p zero
|
|
<*> dtightenS p succ
|
|
dtightenE p (CaseBox qty box ret body) =
|
|
CaseBox qty <$> dtightenE p box
|
|
<*> dtightenS p ret
|
|
<*> dtightenS p body
|
|
dtightenE p (fun :% arg) = [|dtightenE p fun :% tighten p arg|]
|
|
dtightenE p (tm :# ty) = [|dtightenT p tm :# dtightenT p ty|]
|
|
dtightenE p (Coe ty q0 q1 val) =
|
|
[|Coe (dtightenDS p ty) (tighten p q0) (tighten p q1) (dtightenT p val)|]
|
|
dtightenE p (Comp ty q0 q1 val r zero one) =
|
|
[|Comp (dtightenT p ty) (tighten p q0) (tighten p q1)
|
|
(dtightenT p val) (tighten p r)
|
|
(dtightenDS p zero) (dtightenDS p one)|]
|
|
dtightenE p (TypeCase ty ret arms def) =
|
|
[|TypeCase (dtightenE p ty) (dtightenT p ret)
|
|
(traverse (dtightenS p) arms) (dtightenT p def)|]
|
|
dtightenE p (CloE el th) = do
|
|
el <- dtightenE p el
|
|
th <- assert_total $ traverse (dtightenE p) th
|
|
pure $ CloE el th
|
|
dtightenE p (DCloE el th) = do
|
|
th <- tighten p th
|
|
pure $ DCloE el th
|
|
|
|
export
|
|
dtightenS : OPE d1 d2 -> ScopeTermN s d2 n -> Maybe (ScopeTermN s d1 n)
|
|
dtightenS = assert_total $ tightenDScope dtightenT {f = Term}
|
|
|
|
export
|
|
dtightenDS : {s : Nat} -> OPE d1 d2 ->
|
|
DScopeTermN s d2 n -> Maybe (DScopeTermN s d1 n)
|
|
dtightenDS = assert_total $ tightenScope dtightenT
|
|
|
|
|
|
export [TermD] Tighten (\d => Term d n) where tighten p t = dtightenT p t
|
|
export [ElimD] Tighten (\d => Elim d n) where tighten p e = dtightenE p e
|
|
|
|
|
|
-- versions of SY, etc, that try to tighten and use SN automatically
|
|
|
|
public export
|
|
ST : Tighten f => {s : Nat} -> NContext s -> f (s + n) -> Scoped s f n
|
|
ST names body =
|
|
case tightenN s body of
|
|
Just body => S names $ N body
|
|
Nothing => S names $ Y body
|
|
|
|
public export
|
|
DST : {s : Nat} -> NContext s -> Term (s + d) n -> DScopeTermN s d n
|
|
DST names body =
|
|
case tightenN @{TermD} s body of
|
|
Just body => S names $ N body
|
|
Nothing => S names $ Y body
|
|
|
|
public export %inline
|
|
PiT : (qty : Qty) -> (x : BaseName) ->
|
|
(arg : Term d n) -> (res : Term d (S n)) -> Term d n
|
|
PiT {qty, x, arg, res} = Pi {qty, arg, res = ST [< x] res}
|
|
|
|
public export %inline
|
|
SigT : (x : BaseName) -> (fst : Term d n) ->
|
|
(snd : Term d (S n)) -> Term d n
|
|
SigT {x, fst, snd} = Sig {fst, snd = ST [< x] snd}
|
|
|
|
public export %inline
|
|
EqT : (i : BaseName) -> (ty : Term (S d) n) ->
|
|
(l, r : Term d n) -> Term d n
|
|
EqT {i, ty, l, r} = Eq {ty = DST [< i] ty, l, r}
|
|
|
|
public export %inline
|
|
CoeT : (i : BaseName) -> (ty : Term (S d) n) ->
|
|
(p, q : Dim d) -> (val : Term d n) -> Elim d n
|
|
CoeT {i, ty, p, q, val} = Coe {ty = DST [< i] ty, p, q, val}
|
|
|
|
public export %inline
|
|
typeCase1T : Elim d n -> Term d n ->
|
|
(k : TyConKind) -> NContext (arity k) -> Term d (arity k + n) ->
|
|
{default Nat def : Term d n} ->
|
|
Elim d n
|
|
typeCase1T ty ret k ns body {def} =
|
|
typeCase ty ret [(k ** ST ns body)] def
|
|
|
|
|
|
export
|
|
squeeze : {s : Nat} -> ScopeTermN s d n -> ScopeTermN s d n
|
|
squeeze (S names (Y body)) = S names $ maybe (Y body) N $ tightenN s body
|
|
squeeze (S names (N body)) = S names $ N body
|
|
|
|
export
|
|
dsqueeze : {s : Nat} -> DScopeTermN s d n -> DScopeTermN s d n
|
|
dsqueeze (S names (Y body)) =
|
|
S names $ maybe (Y body) N $ tightenN s body @{TermD}
|
|
dsqueeze (S names (N body)) = S names $ N body
|