refactor core syntax slightly to derive Eq/Show

add a new `WithSubst tm env to` record that packages a `tm from`
with a `Subst env from to`, and write instances for just that. the
rest of the AST can be derived
This commit is contained in:
rhiannon morris 2023-04-27 21:37:20 +02:00
parent 7e079a9668
commit 30fa93ab4e
13 changed files with 184 additions and 269 deletions

View file

@ -537,8 +537,8 @@ mutual
Right nt => pure $ Element (TypeCase ty ret arms def) $
tynf `orNo` retnf `orNo` nt
whnf defs ctx (CloE el th) = whnf defs ctx $ pushSubstsWith' id th el
whnf defs ctx (DCloE el th) = whnf defs ctx $ pushSubstsWith' th id el
whnf defs ctx (CloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' id th el
whnf defs ctx (DCloE (Sub el th)) = whnf defs ctx $ pushSubstsWith' th id el
export covering
Whnf Term Reduce.isRedexT where
@ -564,8 +564,8 @@ mutual
Left _ => let tm :# _ = e in pure $ Element tm $ noOr1 $ noOr2 enf
Right na => pure $ Element (E e) $ na `orNo` enf
whnf defs ctx (CloT tm th) = whnf defs ctx $ pushSubstsWith' id th tm
whnf defs ctx (DCloT tm th) = whnf defs ctx $ pushSubstsWith' th id tm
whnf defs ctx (CloT (Sub tm th)) = whnf defs ctx $ pushSubstsWith' id th tm
whnf defs ctx (DCloT (Sub tm th)) = whnf defs ctx $ pushSubstsWith' th id tm
||| reduce a type-case applied to a type constructor
private covering

View file

@ -50,6 +50,13 @@ toEqv Refl {by = SZ} = EqSZ
toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl
export
eqLen : Shift from1 to -> Shift from2 to -> Maybe (from1 = from2)
eqLen SZ SZ = Just Refl
eqLen SZ (SS by) = Nothing
eqLen (SS by) SZ = Nothing
eqLen (SS by) (SS bz) = eqLen by bz
export
0 shiftDiff : (by : Shift from to) -> to = by.nat + from
shiftDiff SZ = Refl

View file

@ -8,8 +8,10 @@ import Quox.Pretty
import Data.Nat
import Data.List
import Data.SnocVect
import Derive.Prelude
%default total
%language ElabReflection
public export
@ -35,8 +37,9 @@ repr (Shift by) = ([], by.nat)
repr (t ::: th) = let (ts, i) = repr th in (t::ts, i)
export Eq (f to) => Eq (Subst f from to) where (==) = (==) `on` repr
export Ord (f to) => Ord (Subst f from to) where compare = compare `on` repr
export Eq (f to) => Eq (Subst f from to) where (==) = (==) `on` repr
export Ord (f to) => Ord (Subst f from to) where compare = compare `on` repr
export Show (f to) => Show (Subst f from to) where show = show . repr
infixl 8 //
@ -155,3 +158,30 @@ prettySubstM pr names bnd op cl th =
export
PrettyHL (f to) => PrettyHL (Subst f from to) where
prettyM th = prettySubstM prettyM (!ask).tnames TVar "[" "]" th
export
eqShape : Subst env from1 to -> Subst env from2 to -> Maybe (from1 = from2)
eqShape (Shift by) (Shift bz) = eqLen by bz
eqShape (Shift by) (t ::: th) = Nothing
eqShape (t ::: th) (Shift by) = Nothing
eqShape (t ::: th) (x ::: ph) = cong S <$> eqShape th ph
public export
record WithSubst tm env n where
constructor Sub
term : tm from
subst : Lazy (Subst env from n)
export
(forall n. Eq (tm n), Eq (env n)) => Eq (WithSubst tm env n) where
Sub t1 s1 == Sub t2 s2 =
case eqShape s1 s2 of
Just Refl => t1 == t2 && s1 == s2
Nothing => False
export %hint
ShowWithSubst : (forall n. Show (tm n), Show (env n)) =>
Show (WithSubst tm env n)
ShowWithSubst = deriveShow

View file

@ -20,8 +20,12 @@ import Data.String
import public Data.SortedMap
import public Data.SortedMap.Dependent
import public Data.SortedSet
import Derive.Prelude
%default total
%language ElabReflection
%hide TT.Name
public export
@ -41,6 +45,37 @@ TagVal : Type
TagVal = String
public export
data ScopedBody : Nat -> (Nat -> Type) -> Nat -> Type where
Y : (body : f (s + n)) -> ScopedBody s f n
N : (body : f n) -> ScopedBody s f n
%name ScopedBody body
export %inline %hint
EqScopedBody : (forall n. Eq (f n)) => Eq (ScopedBody s f n)
EqScopedBody = deriveEq
export %inline %hint
ShowScopedBody : (forall n. Show (f n)) => Show (ScopedBody s f n)
ShowScopedBody = deriveShow
||| a scoped term with names
public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : NContext s
body : ScopedBody s f n
%name Scoped body
export %inline
(forall n. Eq (f n)) => Eq (Scoped s f n) where
s == t = s.body == t.body
export %inline %hint
ShowScoped : (forall n. Show (f n)) => Show (Scoped s f n)
ShowScoped = deriveShow
infixl 8 :#
infixl 9 :@, :%
mutual
@ -51,7 +86,7 @@ mutual
||| first argument `d` is dimension scope size;
||| second `n` is term scope size
public export
data Term : TermLike where
data Term : (d, n : Nat) -> Type where
||| type of types
TYPE : (l : Universe) -> Term d n
@ -90,15 +125,14 @@ mutual
E : (e : Elim d n) -> Term d n
||| term closure/suspended substitution
CloT : (tm : Term d from) -> (th : Lazy (TSubst d from to)) ->
Term d to
CloT : WithSubst (Term d) (Elim d) n -> Term d n
||| dimension closure/suspended substitution
DCloT : (tm : Term dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
Term dto n
DCloT : WithSubst (\d => Term d n) Dim d -> Term d n
%name Term s, t, r
||| first argument `d` is dimension scope size, second `n` is term scope size
public export
data Elim : TermLike where
data Elim : (d, n : Nat) -> Type where
||| free variable
F : (x : Name) -> Elim d n
||| bound variable
@ -157,11 +191,10 @@ mutual
Elim d n
||| term closure/suspended substitution
CloE : (el : Elim d from) -> (th : Lazy (TSubst d from to)) ->
Elim d to
CloE : WithSubst (Elim d) (Elim d) n -> Elim d n
||| dimension closure/suspended substitution
DCloE : (el : Elim dfrom n) -> (th : Lazy (DSubst dfrom dto)) ->
Elim dto n
DCloE : WithSubst (\d => Elim d n) Dim d -> Elim d n
%name Elim e, f
public export
CaseEnumArms : TermLike
@ -180,18 +213,6 @@ mutual
TypeCaseArmBody k = ScopeTermN (arity k)
||| a scoped term with names
public export
record Scoped (s : Nat) (f : Nat -> Type) (n : Nat) where
constructor S
names : NContext s
body : ScopedBody s f n
public export
data ScopedBody : Nat -> (Nat -> Type) -> Nat -> Type where
Y : (body : f (s + n)) -> ScopedBody s f n
N : (body : f n) -> ScopedBody s f n
public export
ScopeTermN, DScopeTermN : Nat -> TermLike
ScopeTermN s d n = Scoped s (Term d) n
@ -202,11 +223,23 @@ mutual
ScopeTerm = ScopeTermN 1
DScopeTerm = DScopeTermN 1
mutual
export %hint
EqTerm : Eq (Term d n)
EqTerm = assert_total {a = Eq (Term d n)} deriveEq
%name Term s, t, r
%name Elim e, f
%name Scoped body
%name ScopedBody body
export %hint
EqElim : Eq (Elim d n)
EqElim = assert_total {a = Eq (Elim d n)} deriveEq
mutual
export %hint
ShowTerm : Show (Term d n)
ShowTerm = assert_total {a = Show (Term d n)} deriveShow
export %hint
ShowElim : Show (Elim d n)
ShowElim = assert_total {a = Show (Elim d n)} deriveShow
||| scope which ignores all its binders
public export %inline

View file

@ -282,14 +282,14 @@ parameters (showSubsts : Bool)
prettyM (E e) = prettyM e
prettyM (CloT s th) =
prettyM (CloT (Sub s th)) =
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM s) <%> prettyTSubst th|]
else
prettyM $ pushSubstsWith' id th s
prettyM (DCloT s th) =
prettyM (DCloT (Sub s th)) =
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM s) <%> prettyDSubst th|]
@ -386,14 +386,14 @@ parameters (showSubsts : Bool)
fromArm (k ** S ns t) =
pure (toSnocList' ns, !(tyCasePat k ns), eterm t.term)
prettyM (CloE e th) =
prettyM (CloE (Sub e th)) =
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyTSubst th|]
else
prettyM $ pushSubstsWith' id th e
prettyM (DCloE e th) =
prettyM (DCloE (Sub e th)) =
if showSubsts then
parensIfM SApp . hang 2 =<<
[|withPrec SApp (prettyM e) <%> prettyDSubst th|]

View file

@ -19,15 +19,15 @@ namespace CanDSubst
||| - otherwise, wraps in a new closure
export
CanDSubst Term where
s // Shift SZ = s
TYPE l // _ = TYPE l
DCloT s ph // th = DCloT s $ ph . th
s // th = DCloT s th
s // Shift SZ = s
TYPE l // _ = TYPE l
DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th
s // th = DCloT $ Sub s th
private
subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n
subDArgs (f :% d) th = subDArgs f th :% (d // th)
subDArgs e th = DCloE e th
subDArgs e th = DCloE $ Sub e th
||| does the minimal reasonable work:
||| - deletes the closure around a term variable
@ -38,12 +38,12 @@ subDArgs e th = DCloE e th
||| - otherwise, wraps in a new closure
export
CanDSubst Elim 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
e // Shift SZ = e
F x // _ = F x
B i // _ = B i
f :% d // th = subDArgs (f :% d) th
DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
e // th = DCloE $ Sub e th
namespace DSubst.ScopeTermN
export %inline
@ -73,12 +73,12 @@ export %inline FromVar (Term d) where fromVar = E . fromVar
||| - otherwise, wraps in a new closure
export
CanSubstSelf (Elim 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
F x // _ = F x
B i // th = th !! i
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
namespace CanTSubst
public export
@ -93,12 +93,12 @@ namespace CanTSubst
||| - 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
TYPE l // _ = TYPE l
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
@ -276,9 +276,9 @@ mutual
pushSubstsWith th ph (Box val) = nclo $ Box $ val // 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 ph (CloT (Sub s ps)) =
pushSubstsWith th (comp th ps ph) s
pushSubstsWith th ph (DCloT s ps) =
pushSubstsWith th ph (DCloT (Sub s ps)) =
pushSubstsWith (ps . th) ph s
export
@ -315,9 +315,9 @@ mutual
pushSubstsWith th ph (TypeCase ty ret arms def) =
nclo $ TypeCase (ty // th // ph) (ret // th // ph)
(map (\t => t // th // ph) arms) (def // th // ph)
pushSubstsWith th ph (CloE e ps) =
pushSubstsWith th ph (CloE (Sub e ps)) =
pushSubstsWith th (comp th ps ph) e
pushSubstsWith th ph (DCloE e ps) =
pushSubstsWith th ph (DCloE (Sub e ps)) =
pushSubstsWith (ps . th) ph e

View file

@ -63,10 +63,12 @@ mutual
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
tightenT p (CloT (Sub 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)|]
pure $ CloT $ Sub tm th
tightenT p (DCloT (Sub tm th)) = do
tm <- tightenT p tm
pure $ DCloT $ Sub tm th
private
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
@ -109,10 +111,12 @@ mutual
<*> tightenT p ret
<*> traverse (tightenS p) arms
<*> tightenT p def
tightenE p (CloE el th) = do
tightenE p (CloE (Sub 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)|]
pure $ CloE $ Sub el th
tightenE p (DCloE (Sub el th)) = do
el <- tightenE p el
pure $ DCloE $ Sub el th
export
tightenS : {s : Nat} -> OPE m n ->
@ -155,11 +159,13 @@ mutual
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
dtightenT p (CloT (Sub 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
pure $ CloT $ Sub tm th
dtightenT p (DCloT (Sub tm th)) = do
th <- tighten p th
pure $ DCloT $ Sub tm th
export
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
@ -195,13 +201,13 @@ mutual
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
dtightenE p (CloE (Sub el th)) = do
el <- dtightenE p el
th <- assert_total $ traverse (dtightenE p) th
pure $ CloE el th
dtightenE p (DCloE el th) = do
pure $ CloE $ Sub el th
dtightenE p (DCloE (Sub el th)) = do
th <- tighten p th
pure $ DCloE el th
pure $ DCloE $ Sub el th
export
dtightenS : OPE d1 d2 -> ScopeTermN s d2 n -> Maybe (ScopeTermN s d1 n)