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

View file

@ -50,6 +50,13 @@ toEqv Refl {by = SZ} = EqSZ
toEqv Refl {by = (SS by)} = EqSS $ toEqv Refl 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 export
0 shiftDiff : (by : Shift from to) -> to = by.nat + from 0 shiftDiff : (by : Shift from to) -> to = by.nat + from
shiftDiff SZ = Refl shiftDiff SZ = Refl

View file

@ -8,8 +8,10 @@ import Quox.Pretty
import Data.Nat import Data.Nat
import Data.List import Data.List
import Data.SnocVect import Data.SnocVect
import Derive.Prelude
%default total %default total
%language ElabReflection
public export public export
@ -35,8 +37,9 @@ repr (Shift by) = ([], by.nat)
repr (t ::: th) = let (ts, i) = repr th in (t::ts, i) 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 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 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 // infixl 8 //
@ -155,3 +158,30 @@ prettySubstM pr names bnd op cl th =
export export
PrettyHL (f to) => PrettyHL (Subst f from to) where PrettyHL (f to) => PrettyHL (Subst f from to) where
prettyM th = prettySubstM prettyM (!ask).tnames TVar "[" "]" th 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
import public Data.SortedMap.Dependent import public Data.SortedMap.Dependent
import public Data.SortedSet import public Data.SortedSet
import Derive.Prelude
%default total %default total
%language ElabReflection
%hide TT.Name
public export public export
@ -41,6 +45,37 @@ TagVal : Type
TagVal = String 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 8 :#
infixl 9 :@, :% infixl 9 :@, :%
mutual mutual
@ -51,7 +86,7 @@ mutual
||| first argument `d` is dimension scope size; ||| first argument `d` is dimension scope size;
||| second `n` is term scope size ||| second `n` is term scope size
public export public export
data Term : TermLike where data Term : (d, n : Nat) -> Type where
||| type of types ||| type of types
TYPE : (l : Universe) -> Term d n TYPE : (l : Universe) -> Term d n
@ -90,15 +125,14 @@ mutual
E : (e : Elim d n) -> Term d n E : (e : Elim d n) -> Term d n
||| term closure/suspended substitution ||| term closure/suspended substitution
CloT : (tm : Term d from) -> (th : Lazy (TSubst d from to)) -> CloT : WithSubst (Term d) (Elim d) n -> Term d n
Term d to
||| dimension closure/suspended substitution ||| dimension closure/suspended substitution
DCloT : (tm : Term dfrom n) -> (th : Lazy (DSubst dfrom dto)) -> DCloT : WithSubst (\d => Term d n) Dim d -> Term d n
Term dto n %name Term s, t, r
||| first argument `d` is dimension scope size, second `n` is term scope size ||| first argument `d` is dimension scope size, second `n` is term scope size
public export public export
data Elim : TermLike where data Elim : (d, n : Nat) -> Type where
||| free variable ||| free variable
F : (x : Name) -> Elim d n F : (x : Name) -> Elim d n
||| bound variable ||| bound variable
@ -157,11 +191,10 @@ mutual
Elim d n Elim d n
||| term closure/suspended substitution ||| term closure/suspended substitution
CloE : (el : Elim d from) -> (th : Lazy (TSubst d from to)) -> CloE : WithSubst (Elim d) (Elim d) n -> Elim d n
Elim d to
||| dimension closure/suspended substitution ||| dimension closure/suspended substitution
DCloE : (el : Elim dfrom n) -> (th : Lazy (DSubst dfrom dto)) -> DCloE : WithSubst (\d => Elim d n) Dim d -> Elim d n
Elim dto n %name Elim e, f
public export public export
CaseEnumArms : TermLike CaseEnumArms : TermLike
@ -180,18 +213,6 @@ mutual
TypeCaseArmBody k = ScopeTermN (arity k) 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 public export
ScopeTermN, DScopeTermN : Nat -> TermLike ScopeTermN, DScopeTermN : Nat -> TermLike
ScopeTermN s d n = Scoped s (Term d) n ScopeTermN s d n = Scoped s (Term d) n
@ -202,11 +223,23 @@ mutual
ScopeTerm = ScopeTermN 1 ScopeTerm = ScopeTermN 1
DScopeTerm = DScopeTermN 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 export %hint
%name Elim e, f EqElim : Eq (Elim d n)
%name Scoped body EqElim = assert_total {a = Eq (Elim d n)} deriveEq
%name ScopedBody body
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 ||| scope which ignores all its binders
public export %inline public export %inline

View file

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

View file

@ -19,15 +19,15 @@ namespace CanDSubst
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanDSubst Term where CanDSubst Term where
s // Shift SZ = s s // Shift SZ = s
TYPE l // _ = TYPE l TYPE l // _ = TYPE l
DCloT s ph // th = DCloT s $ ph . th DCloT (Sub s ph) // th = DCloT $ Sub s $ ph . th
s // th = DCloT s th s // th = DCloT $ Sub s th
private private
subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n subDArgs : Elim dfrom n -> DSubst dfrom dto -> Elim dto n
subDArgs (f :% d) th = subDArgs f th :% (d // th) 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: ||| does the minimal reasonable work:
||| - deletes the closure around a term variable ||| - deletes the closure around a term variable
@ -38,12 +38,12 @@ subDArgs e th = DCloE e th
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanDSubst Elim where CanDSubst Elim where
e // Shift SZ = e e // Shift SZ = e
F x // _ = F x F x // _ = F x
B i // _ = B i B i // _ = B i
f :% d // th = subDArgs (f :% d) th f :% d // th = subDArgs (f :% d) th
DCloE e ph // th = DCloE e $ ph . th DCloE (Sub e ph) // th = DCloE $ Sub e $ ph . th
e // th = DCloE e th e // th = DCloE $ Sub e th
namespace DSubst.ScopeTermN namespace DSubst.ScopeTermN
export %inline export %inline
@ -73,12 +73,12 @@ export %inline FromVar (Term d) where fromVar = E . fromVar
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanSubstSelf (Elim d) where CanSubstSelf (Elim d) where
F x // _ = F x F x // _ = F x
B i // th = th !! i B i // th = th !! i
CloE e ph // th = assert_total CloE e $ ph . th CloE (Sub e ph) // th = assert_total CloE $ Sub e $ ph . th
e // th = case force th of e // th = case force th of
Shift SZ => e Shift SZ => e
th => CloE e th th => CloE $ Sub e th
namespace CanTSubst namespace CanTSubst
public export public export
@ -93,12 +93,12 @@ namespace CanTSubst
||| - otherwise, wraps in a new closure ||| - otherwise, wraps in a new closure
export export
CanTSubst Term where CanTSubst Term where
TYPE l // _ = TYPE l TYPE l // _ = TYPE l
E e // th = E $ e // th E e // th = E $ e // th
CloT s ph // th = CloT s $ ph . th CloT (Sub s ph) // th = CloT $ Sub s $ ph . th
s // th = case force th of s // th = case force th of
Shift SZ => s Shift SZ => s
th => CloT s th th => CloT $ Sub s th
namespace ScopeTermN namespace ScopeTermN
export %inline export %inline
@ -276,9 +276,9 @@ mutual
pushSubstsWith th ph (Box val) = nclo $ Box $ val // th // ph pushSubstsWith th ph (Box val) = nclo $ Box $ val // th // ph
pushSubstsWith th ph (E e) = pushSubstsWith th ph (E e) =
let Element e nc = pushSubstsWith th ph e in nclo $ 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 (comp th ps ph) s
pushSubstsWith th ph (DCloT s ps) = pushSubstsWith th ph (DCloT (Sub s ps)) =
pushSubstsWith (ps . th) ph s pushSubstsWith (ps . th) ph s
export export
@ -315,9 +315,9 @@ mutual
pushSubstsWith th ph (TypeCase ty ret arms def) = pushSubstsWith th ph (TypeCase ty ret arms def) =
nclo $ TypeCase (ty // th // ph) (ret // th // ph) nclo $ TypeCase (ty // th // ph) (ret // th // ph)
(map (\t => t // th // ph) arms) (def // 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 (comp th ps ph) e
pushSubstsWith th ph (DCloE e ps) = pushSubstsWith th ph (DCloE (Sub e ps)) =
pushSubstsWith (ps . th) ph e 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 qty ty) = BOX qty <$> tightenT p ty
tightenT p (Box val) = Box <$> tightenT p val tightenT p (Box val) = Box <$> tightenT p val
tightenT p (E e) = assert_total $ E <$> tightenE p e 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 th <- assert_total $ tightenSub tightenE p th
pure $ CloT tm th pure $ CloT $ Sub tm th
tightenT p (DCloT tm th) = [|DCloT (tightenT p tm) (pure th)|] tightenT p (DCloT (Sub tm th)) = do
tm <- tightenT p tm
pure $ DCloT $ Sub tm th
private private
tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1) tightenE : OPE n1 n2 -> Elim d n2 -> Maybe (Elim d n1)
@ -109,10 +111,12 @@ mutual
<*> tightenT p ret <*> tightenT p ret
<*> traverse (tightenS p) arms <*> traverse (tightenS p) arms
<*> tightenT p def <*> tightenT p def
tightenE p (CloE el th) = do tightenE p (CloE (Sub el th)) = do
th <- assert_total $ tightenSub tightenE p th th <- assert_total $ tightenSub tightenE p th
pure $ CloE el th pure $ CloE $ Sub el th
tightenE p (DCloE el th) = [|DCloE (tightenE p el) (pure th)|] tightenE p (DCloE (Sub el th)) = do
el <- tightenE p el
pure $ DCloE $ Sub el th
export export
tightenS : {s : Nat} -> OPE m n -> tightenS : {s : Nat} -> OPE m n ->
@ -155,11 +159,13 @@ mutual
dtightenT p (BOX qty ty) = BOX qty <$> dtightenT p ty dtightenT p (BOX qty ty) = BOX qty <$> dtightenT p ty
dtightenT p (Box val) = Box <$> dtightenT p val dtightenT p (Box val) = Box <$> dtightenT p val
dtightenT p (E e) = assert_total $ E <$> dtightenE p e 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 tm <- dtightenT p tm
th <- assert_total $ traverse (dtightenE p) th th <- assert_total $ traverse (dtightenE p) th
pure $ CloT tm th pure $ CloT $ Sub tm th
dtightenT p (DCloT tm th) = do th <- tighten p th; pure $ DCloT tm th dtightenT p (DCloT (Sub tm th)) = do
th <- tighten p th
pure $ DCloT $ Sub tm th
export export
dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n) dtightenE : OPE d1 d2 -> Elim d2 n -> Maybe (Elim d1 n)
@ -195,13 +201,13 @@ mutual
dtightenE p (TypeCase ty ret arms def) = dtightenE p (TypeCase ty ret arms def) =
[|TypeCase (dtightenE p ty) (dtightenT p ret) [|TypeCase (dtightenE p ty) (dtightenT p ret)
(traverse (dtightenS p) arms) (dtightenT p def)|] (traverse (dtightenS p) arms) (dtightenT p def)|]
dtightenE p (CloE el th) = do dtightenE p (CloE (Sub el th)) = do
el <- dtightenE p el el <- dtightenE p el
th <- assert_total $ traverse (dtightenE p) th th <- assert_total $ traverse (dtightenE p) th
pure $ CloE el th pure $ CloE $ Sub el th
dtightenE p (DCloE el th) = do dtightenE p (DCloE (Sub el th)) = do
th <- tighten p th th <- tighten p th
pure $ DCloE el th pure $ DCloE $ Sub el th
export export
dtightenS : OPE d1 d2 -> ScopeTermN s d2 n -> Maybe (ScopeTermN s d1 n) dtightenS : OPE d1 d2 -> ScopeTermN s d2 n -> Maybe (ScopeTermN s d1 n)

View file

@ -1,158 +0,0 @@
module TermImpls
import Quox.Syntax
import public Quox.Pretty
private
eqShiftLen : Shift from1 to -> Shift from2 to -> Maybe (from1 = from2)
eqShiftLen SZ SZ = Just Refl
eqShiftLen (SS by) (SS bz) = eqShiftLen by bz
eqShiftLen _ _ = Nothing
private
eqSubstLen : Subst tm1 from1 to -> Subst tm2 from2 to -> Maybe (from1 = from2)
eqSubstLen (Shift by) (Shift bz) = eqShiftLen by bz
eqSubstLen (_ ::: th) (_ ::: ph) = cong S <$> eqSubstLen th ph
eqSubstLen _ _ = Nothing
-- maybe from1 = from2 in the last case, but this is for
-- (==), and the substs aren't equal, so who cares
mutual
export covering
Eq (Term d n) where
TYPE k == TYPE l = k == l
TYPE _ == _ = False
Pi qty1 arg1 res1 == Pi qty2 arg2 res2 =
qty1 == qty2 && arg1 == arg2 && res1 == res2
Pi {} == _ = False
Lam body1 == Lam body2 = body1 == body2
Lam {} == _ = False
Sig fst1 snd1 == Sig fst2 snd2 =
fst1 == fst2 && snd1 == snd2
Sig {} == _ = False
Pair fst1 snd1 == Pair fst2 snd2 = fst1 == fst2 && snd1 == snd2
Pair {} == _ = False
Enum ts1 == Enum ts2 = ts1 == ts2
Enum _ == _ = False
Tag t1 == Tag t2 = t1 == t2
Tag _ == _ = False
Eq ty1 l1 r1 == Eq ty2 l2 r2 =
ty1 == ty2 && l1 == l2 && r1 == r2
Eq {} == _ = False
DLam body1 == DLam body2 = body1 == body2
DLam {} == _ = False
Nat == Nat = True
Nat == _ = False
Zero == Zero = True
Zero == _ = False
Succ m == Succ n = m == n
Succ _ == _ = False
BOX q1 ty1 == BOX q2 ty2 = q1 == q2 && ty1 == ty2
BOX {} == _ = False
Box val1 == Box val2 = val1 == val2
Box _ == _ = False
E e == E f = e == f
E _ == _ = False
CloT tm1 th1 == CloT tm2 th2 =
case eqSubstLen th1 th2 of
Just Refl => tm1 == tm2 && th1 == th2
Nothing => False
CloT {} == _ = False
DCloT tm1 th1 == DCloT tm2 th2 =
case eqSubstLen th1 th2 of
Just Refl => tm1 == tm2 && th1 == th2
Nothing => False
DCloT {} == _ = False
export covering
Eq (Elim d n) where
F x == F y = x == y
F _ == _ = False
B i == B j = i == j
B _ == _ = False
(fun1 :@ arg1) == (fun2 :@ arg2) = fun1 == fun2 && arg1 == arg2
(_ :@ _) == _ = False
CasePair q1 p1 r1 b1 == CasePair q2 p2 r2 b2 =
q1 == q2 && p1 == p2 && r1 == r2 && b1 == b2
CasePair {} == _ = False
CaseEnum q1 t1 r1 a1 == CaseEnum q2 t2 r2 a2 =
q1 == q2 && t1 == t2 && r1 == r2 && a1 == a2
CaseEnum {} == _ = False
CaseNat q1 q1' n1 r1 z1 s1 == CaseNat q2 q2' n2 r2 z2 s2 =
q1 == q2 && q1' == q2' && n1 == n2 &&
r1 == r2 && z1 == z2 && s1 == s2
CaseNat {} == _ = False
CaseBox q1 x1 r1 b1 == CaseBox q2 x2 r2 b2 =
q1 == q2 && x1 == x2 && r1 == r2 && b1 == b2
CaseBox {} == _ = False
(fun1 :% dim1) == (fun2 :% dim2) = fun1 == fun2 && dim1 == dim2
(_ :% _) == _ = False
(tm1 :# ty1) == (tm2 :# ty2) = tm1 == tm2 && ty1 == ty2
(_ :# _) == _ = False
Coe ty1 p1 q1 val1 == Coe ty2 p2 q2 val2 =
ty1 == ty2 && p1 == p2 && q1 == q2 && val1 == val2
Coe {} == _ = False
Comp ty1 p1 q1 val1 r1 zero1 one1 == Comp ty2 p2 q2 val2 r2 zero2 one2 =
ty1 == ty2 && p1 == p2 && q1 == q2 &&
val1 == val2 && r1 == r2 && zero1 == zero2 && one1 == one2
Comp {} == _ = False
TypeCase ty1 ret1 arms1 def1 == TypeCase ty2 ret2 arms2 def2 =
ty1 == ty2 && ret1 == ret2 &&
arms1 == arms2 && def1 == def2
TypeCase {} == _ = False
CloE el1 th1 == CloE el2 th2 =
case eqSubstLen th1 th2 of
Just Refl => el1 == el2 && th1 == th2
Nothing => False
CloE {} == _ = False
DCloE el1 th1 == DCloE el2 th2 =
case eqSubstLen th1 th2 of
Just Refl => el1 == el2 && th1 == th2
Nothing => False
DCloE {} == _ = False
export covering
{s : Nat} -> Eq (ScopeTermN s d n) where
b1 == b2 = b1.term == b2.term
export covering
{s : Nat} -> Eq (DScopeTermN s d n) where
b1 == b2 = b1.term == b2.term
export covering
Show (Term d n) where
showPrec d t = showParens (d /= Open) $ prettyStr True t
export covering
Show (Elim d n) where
showPrec d e = showParens (d /= Open) $ prettyStr True e

View file

@ -237,27 +237,27 @@ tests = "equality & subtyping" :- [
testEq "[#0]{} = [#0] : A" $ testEq "[#0]{} = [#0] : A" $
equalT (extendTy Any "x" (FT "A") empty) equalT (extendTy Any "x" (FT "A") empty)
(FT "A") (FT "A")
(CloT (BVT 0) id) (CloT (Sub (BVT 0) id))
(BVT 0), (BVT 0),
testEq "[#0]{a} = [a] : A" $ testEq "[#0]{a} = [a] : A" $
equalT empty (FT "A") equalT empty (FT "A")
(CloT (BVT 0) (F "a" ::: id)) (CloT (Sub (BVT 0) (F "a" ::: id)))
(FT "a"), (FT "a"),
testEq "[#0]{a,b} = [a] : A" $ testEq "[#0]{a,b} = [a] : A" $
equalT empty (FT "A") equalT empty (FT "A")
(CloT (BVT 0) (F "a" ::: F "b" ::: id)) (CloT (Sub (BVT 0) (F "a" ::: F "b" ::: id)))
(FT "a"), (FT "a"),
testEq "[#1]{a,b} = [b] : A" $ testEq "[#1]{a,b} = [b] : A" $
equalT empty (FT "A") equalT empty (FT "A")
(CloT (BVT 1) (F "a" ::: F "b" ::: id)) (CloT (Sub (BVT 1) (F "a" ::: F "b" ::: id)))
(FT "b"), (FT "b"),
testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (N)" $ testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (N)" $
equalT empty (Arr Zero (FT "B") (FT "A")) equalT empty (Arr Zero (FT "B") (FT "A"))
(CloT (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id)) (CloT (Sub (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id)))
(Lam $ S [< "y"] $ N $ FT "a"), (Lam $ S [< "y"] $ N $ FT "a"),
testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (Y)" $ testEq "(λy ⇒ [#1]){a} = λy ⇒ [a] : B ⇾ A (Y)" $
equalT empty (Arr Zero (FT "B") (FT "A")) equalT empty (Arr Zero (FT "B") (FT "A"))
(CloT ([< "y"] :\\ BVT 1) (F "a" ::: id)) (CloT (Sub ([< "y"] :\\ BVT 1) (F "a" ::: id)))
([< "y"] :\\ FT "a") ([< "y"] :\\ FT "a")
], ],
@ -265,12 +265,12 @@ tests = "equality & subtyping" :- [
testEq "★₀‹𝟎› = ★₀ : ★₁" $ testEq "★₀‹𝟎› = ★₀ : ★₁" $
equalTD 1 equalTD 1
(extendDim "𝑗" empty) (extendDim "𝑗" empty)
(TYPE 1) (DCloT (TYPE 0) (K Zero ::: id)) (TYPE 0), (TYPE 1) (DCloT (Sub (TYPE 0) (K Zero ::: id))) (TYPE 0),
testEq "(δ i ⇒ a)𝟎 = (δ i ⇒ a) : (a ≡ a : A)" $ testEq "(δ i ⇒ a)𝟎 = (δ i ⇒ a) : (a ≡ a : A)" $
equalTD 1 equalTD 1
(extendDim "𝑗" empty) (extendDim "𝑗" empty)
(Eq0 (FT "A") (FT "a") (FT "a")) (Eq0 (FT "A") (FT "a") (FT "a"))
(DCloT ([< "i"] :\\% FT "a") (K Zero ::: id)) (DCloT (Sub ([< "i"] :\\% FT "a") (K Zero ::: id)))
([< "i"] :\\% FT "a"), ([< "i"] :\\% FT "a"),
note "it is hard to think of well-typed terms with big dctxs" note "it is hard to think of well-typed terms with big dctxs"
], ],
@ -504,10 +504,10 @@ tests = "equality & subtyping" :- [
"elim closure" :- [ "elim closure" :- [
testEq "#0{a} = a" $ testEq "#0{a} = a" $
equalE empty (CloE (BV 0) (F "a" ::: id)) (F "a"), equalE empty (CloE (Sub (BV 0) (F "a" ::: id))) (F "a"),
testEq "#1{a} = #0" $ testEq "#1{a} = #0" $
equalE (extendTy Any "x" (FT "A") empty) equalE (extendTy Any "x" (FT "A") empty)
(CloE (BV 1) (F "a" ::: id)) (BV 0) (CloE (Sub (BV 1) (F "a" ::: id))) (BV 0)
], ],
"elim d-closure" :- [ "elim d-closure" :- [
@ -515,41 +515,42 @@ tests = "equality & subtyping" :- [
testEq "(eq-AB #0)𝟎 = eq-AB 𝟎" $ testEq "(eq-AB #0)𝟎 = eq-AB 𝟎" $
equalED 1 equalED 1
(extendDim "𝑖" empty) (extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (K Zero ::: id)) (DCloE (Sub (F "eq-AB" :% BV 0) (K Zero ::: id)))
(F "eq-AB" :% K Zero), (F "eq-AB" :% K Zero),
testEq "(eq-AB #0)𝟎 = A" $ testEq "(eq-AB #0)𝟎 = A" $
equalED 1 equalED 1
(extendDim "𝑖" empty) (extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (K Zero ::: id)) (F "A"), (DCloE (Sub (F "eq-AB" :% BV 0) (K Zero ::: id))) (F "A"),
testEq "(eq-AB #0)𝟏 = B" $ testEq "(eq-AB #0)𝟏 = B" $
equalED 1 equalED 1
(extendDim "𝑖" empty) (extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (K One ::: id)) (F "B"), (DCloE (Sub (F "eq-AB" :% BV 0) (K One ::: id))) (F "B"),
testNeq "(eq-AB #0)𝟏 ≠ A" $ testNeq "(eq-AB #0)𝟏 ≠ A" $
equalED 1 equalED 1
(extendDim "𝑖" empty) (extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (K One ::: id)) (F "A"), (DCloE (Sub (F "eq-AB" :% BV 0) (K One ::: id))) (F "A"),
testEq "(eq-AB #0)#0,𝟎 = (eq-AB #0)" $ testEq "(eq-AB #0)#0,𝟎 = (eq-AB #0)" $
equalED 2 equalED 2
(extendDim "𝑗" $ extendDim "𝑖" empty) (extendDim "𝑗" $ extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id)) (DCloE (Sub (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id)))
(F "eq-AB" :% BV 0), (F "eq-AB" :% BV 0),
testNeq "(eq-AB #0)𝟎 ≠ (eq-AB 𝟎)" $ testNeq "(eq-AB #0)𝟎 ≠ (eq-AB 𝟎)" $
equalED 2 equalED 2
(extendDim "𝑗" $ extendDim "𝑖" empty) (extendDim "𝑗" $ extendDim "𝑖" empty)
(DCloE (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id)) (DCloE (Sub (F "eq-AB" :% BV 0) (BV 0 ::: K Zero ::: id)))
(F "eq-AB" :% K Zero), (F "eq-AB" :% K Zero),
testEq "#0𝟎 = #0 # term and dim vars distinct" $ testEq "#0𝟎 = #0 # term and dim vars distinct" $
equalED 1 equalED 1
(extendTy Any "x" (FT "A") $ extendDim "𝑖" empty) (extendTy Any "x" (FT "A") $ extendDim "𝑖" empty)
(DCloE (BV 0) (K Zero ::: id)) (BV 0), (DCloE (Sub (BV 0) (K Zero ::: id))) (BV 0),
testEq "a𝟎 = a" $ testEq "a𝟎 = a" $
equalED 1 (extendDim "𝑖" empty) (DCloE (F "a") (K Zero ::: id)) (F "a"), equalED 1 (extendDim "𝑖" empty)
(DCloE (Sub (F "a") (K Zero ::: id))) (F "a"),
testEq "(f [a])𝟎 = f𝟎 [a]𝟎" $ testEq "(f [a])𝟎 = f𝟎 [a]𝟎" $
let th = K Zero ::: id in let th = K Zero ::: id in
equalED 1 (extendDim "𝑖" empty) equalED 1 (extendDim "𝑖" empty)
(DCloE (F "f" :@ FT "a") th) (DCloE (Sub (F "f" :@ FT "a") th))
(DCloE (F "f") th :@ DCloT (FT "a") th) (DCloE (Sub (F "f") th) :@ DCloT (Sub (FT "a") th))
], ],
"clashes" :- [ "clashes" :- [

View file

@ -2,7 +2,6 @@ module Tests.FromPTerm
import Quox.Parser.FromParser import Quox.Parser.FromParser
import Quox.Parser import Quox.Parser
import TermImpls
import TypingImpls import TypingImpls
import Tests.Parser as TParser import Tests.Parser as TParser
import TAP import TAP

View file

@ -2,7 +2,6 @@ module Tests.Reduce
import Quox.Syntax as Lib import Quox.Syntax as Lib
import Quox.Equal import Quox.Equal
import TermImpls
import TypingImpls import TypingImpls
import TAP import TAP
@ -67,34 +66,34 @@ tests = "whnf" :- [
"elim closure" :- [ "elim closure" :- [
testWhnf "x{}" (ctx [< ("A", Nat)]) testWhnf "x{}" (ctx [< ("A", Nat)])
(CloE (BV 0) id) (CloE (Sub (BV 0) id))
(BV 0), (BV 0),
testWhnf "x{a/x}" empty testWhnf "x{a/x}" empty
(CloE (BV 0) (F "a" ::: id)) (CloE (Sub (BV 0) (F "a" ::: id)))
(F "a"), (F "a"),
testWhnf "x{x/x,a/y}" (ctx [< ("A", Nat)]) testWhnf "x{x/x,a/y}" (ctx [< ("A", Nat)])
(CloE (BV 0) (BV 0 ::: F "a" ::: id)) (CloE (Sub (BV 0) (BV 0 ::: F "a" ::: id)))
(BV 0), (BV 0),
testWhnf "x{(y{a/y})/x}" empty testWhnf "x{(y{a/y})/x}" empty
(CloE (BV 0) ((CloE (BV 0) (F "a" ::: id)) ::: id)) (CloE (Sub (BV 0) ((CloE (Sub (BV 0) (F "a" ::: id))) ::: id)))
(F "a"), (F "a"),
testWhnf "(x y){f/x,a/y}" empty testWhnf "(x y){f/x,a/y}" empty
(CloE (BV 0 :@ BVT 1) (F "f" ::: F "a" ::: id)) (CloE (Sub (BV 0 :@ BVT 1) (F "f" ::: F "a" ::: id)))
(F "f" :@ FT "a"), (F "f" :@ FT "a"),
testWhnf "([y] ∷ [x]){A/x}" (ctx [< ("A", Nat)]) testWhnf "([y] ∷ [x]){A/x}" (ctx [< ("A", Nat)])
(CloE (BVT 1 :# BVT 0) (F "A" ::: id)) (CloE (Sub (BVT 1 :# BVT 0) (F "A" ::: id)))
(BV 0), (BV 0),
testWhnf "([y] ∷ [x]){A/x,a/y}" empty testWhnf "([y] ∷ [x]){A/x,a/y}" empty
(CloE (BVT 1 :# BVT 0) (F "A" ::: F "a" ::: id)) (CloE (Sub (BVT 1 :# BVT 0) (F "A" ::: F "a" ::: id)))
(F "a") (F "a")
], ],
"term closure" :- [ "term closure" :- [
testWhnf "(λy. x){a/x}" empty testWhnf "(λy. x){a/x}" empty
(CloT (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id)) (CloT (Sub (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id)))
(Lam $ S [< "y"] $ N $ FT "a"), (Lam $ S [< "y"] $ N $ FT "a"),
testWhnf "(λy. y){a/x}" empty testWhnf "(λy. y){a/x}" empty
(CloT ([< "y"] :\\ BVT 0) (F "a" ::: id)) (CloT (Sub ([< "y"] :\\ BVT 0) (F "a" ::: id)))
([< "y"] :\\ BVT 0) ([< "y"] :\\ BVT 0)
], ],
@ -112,8 +111,8 @@ tests = "whnf" :- [
F "a" :@ F "a" :@
E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"), E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"),
testNoStep "λx. [y [x]]{x/x,a/y}" (ctx [< ("A", Nat)]) $ testNoStep "λx. [y [x]]{x/x,a/y}" (ctx [< ("A", Nat)]) $
[< "x"] :\\ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id), [< "x"] :\\ CloT (Sub (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id)),
testNoStep "f ([y [x]]{x/x,a/y})" (ctx [< ("A", Nat)]) $ testNoStep "f ([y [x]]{x/x,a/y})" (ctx [< ("A", Nat)]) $
F "f" :@ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id) F "f" :@ CloT (Sub (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id))
] ]
] ]

View file

@ -3,7 +3,6 @@ module TypingImpls
import TAP import TAP
import public Quox.Typing import public Quox.Typing
import public Quox.Pretty import public Quox.Pretty
import public TermImpls
import Derive.Prelude import Derive.Prelude
%language ElabReflection %language ElabReflection

View file

@ -5,7 +5,6 @@ depends = base, contrib, elab-util, snocvect, quox-lib, tap, eff
executable = quox-tests executable = quox-tests
main = Tests main = Tests
modules = modules =
TermImpls,
TypingImpls, TypingImpls,
PrettyExtra, PrettyExtra,
Tests.DimEq, Tests.DimEq,