typed equality

This commit is contained in:
rhiannon morris 2023-02-10 21:40:44 +01:00
parent 3b13f0a82c
commit 42798f243f
8 changed files with 410 additions and 250 deletions

View file

@ -35,6 +35,10 @@ mkAbstract : IsQty q => (qty : q) -> (0 _ : IsGlobal qty) =>
mkAbstract qty type = MkDef' {qty, type = T type, term = Nothing} mkAbstract qty type = MkDef' {qty, type = T type, term = Nothing}
public export %inline
(.get0) : AnyTerm q -> Term q 0 0
t.get0 = t.get
public export %inline public export %inline
(.type0) : Definition' q _ -> Term q 0 0 (.type0) : Definition' q _ -> Term q 0 0
g.type0 = g.type.get g.type0 = g.type.get
@ -117,10 +121,10 @@ NonRedexTerm q d n defs = Subset (Term q d n) (NotRedex defs)
parameters {0 isGlobal : _} (defs : Definitions' q isGlobal) parameters {0 isGlobal : _} (defs : Definitions' q isGlobal)
namespace Term namespace Term
export %inline export %inline
whnf : Term q d n -> NonRedexTerm q d n defs whnfD : Term q d n -> NonRedexTerm q d n defs
whnf = whnf $ \x => lookupElim x defs whnfD = whnf $ \x => lookupElim x defs
namespace Elim namespace Elim
export %inline export %inline
whnf : Elim q d n -> NonRedexElim q d n defs whnfD : Elim q d n -> NonRedexElim q d n defs
whnf = whnf $ \x => lookupElim x defs whnfD = whnf $ \x => lookupElim x defs

View file

@ -9,8 +9,8 @@ import Data.Maybe
private %inline private %inline
ClashE : EqMode -> Elim q d n -> Elim q d n -> Error q ClashE : EqMode -> Term q d n -> Elim q d n -> Elim q d n -> Error q
ClashE mode = ClashT mode `on` E ClashE mode ty = ClashT mode ty `on` E
public export public export
@ -34,179 +34,262 @@ mode : HasEnv m => m EqMode
mode = asks mode mode = asks mode
private %inline private %inline
clashT : CanEqual q m => Term q d n -> Term q d n -> m a clashT : CanEqual q m => Term q d n -> Term q d n -> Term q d n -> m a
clashT s t = throwError $ ClashT !mode s t clashT ty s t = throwError $ ClashT !mode ty s t
private %inline private %inline
clashE : CanEqual q m => Elim q d n -> Elim q d n -> m a clashE : CanEqual q m => Elim q d n -> Elim q d n -> m a
clashE e f = throwError $ ClashE !mode e f clashE e f = throwError $ ClashE !mode e f
parameters {0 isGlobal : _} (defs : Definitions' q isGlobal) public export %inline
isType : (t : Term {}) -> Bool
isType (TYPE {}) = True
isType (Pi {}) = True
isType (Lam {}) = False
isType (Sig {}) = True
isType (Pair {}) = False
isType (Eq {}) = True
isType (DLam {}) = False
isType (E {}) = True
isType (CloT {}) = False
isType (DCloT {}) = False
parameters {auto _ : HasErr q m}
export %inline
ensure : (a -> Error q) -> (p : a -> Bool) -> (t : a) -> m (So (p t))
ensure e p t = case nchoose $ p t of
Left y => pure y
Right n => throwError $ e t
export %inline
ensureType : (t : Term q d n) -> m (So (isType t))
ensureType = ensure NotType isType
parameters (defs : Definitions' q _) {auto _ : (CanEqual q m, Eq q)}
mutual mutual
-- [todo] remove cumulativity & subtyping, it's too much of a pain
-- mugen might be good
namespace Term namespace Term
export covering export covering %inline
compareN' : CanEqual q m => Eq q => compare0 : TContext q 0 n -> (ty, s, t : Term q 0 n) -> m ()
(s, t : Term q 0 n) -> compare0 ctx ty s t = do
(0 _ : NotRedex defs s) -> (0 _ : NotRedex defs t) -> let Element ty nty = whnfD defs ty
Element s ns = whnfD defs s
Element t nt = whnfD defs t
tty <- ensureType ty
compare0' ctx ty s t
private %inline
toLamBody : Elim q d n -> Term q d (S n)
toLamBody e = E $ weakE e :@ BVT 0
private covering
compare0' : TContext q 0 n ->
(ty, s, t : Term q 0 n) ->
(0 nty : NotRedex defs ty) => (0 tty : So (isType ty)) =>
(0 ns : NotRedex defs s) => (0 nt : NotRedex defs t) =>
m () m ()
compare0' ctx (TYPE _) s t = compareType ctx s t
compareN' (TYPE k) (TYPE l) _ _ = expectModeU !mode k l compare0' ctx ty@(Pi {arg, res, _}) s t = local {mode := Equal} $
compareN' s@(TYPE _) t _ _ = clashT s t let ctx' = ctx :< arg
eta : Elim q 0 ? -> ScopeTerm q 0 ? -> m ()
eta e (TUsed b) = compare0 ctx' res.term (toLamBody e) b
eta e (TUnused _) = clashT ty s t
in
case (s, t) of
(Lam _ b1, Lam _ b2) => compare0 ctx' res.term b1.term b2.term
(E e, Lam _ b) => eta e b
(Lam _ b, E e) => eta e b
(E e, E f) => ignore $ compare0 ctx e f
_ => throwError $ WrongType ty s t
compareN' (Pi qty1 _ arg1 res1) (Pi qty2 _ arg2 res2) _ _ = do compare0' ctx ty@(Sig {fst, snd, _}) s t = local {mode := Equal} $
expectEqualQ qty1 qty2 -- no η (no fst/snd for π ≱ 0), so…
compare0 arg2 arg1 -- reversed for contravariant domain -- [todo] η for π ≥ 0 maybe
compare0 res1 res2 case (s, t) of
compareN' s@(Pi {}) t _ _ = clashT s t (Pair sFst sSnd, Pair tFst tSnd) => do
compare0 ctx fst sFst tFst
compare0 ctx (sub1 snd (sFst :# fst)) sSnd tSnd
_ => throwError $ WrongType ty s t
-- [todo] eta -- ✨ uip ✨
compareN' (Lam _ body1) (Lam _ body2) _ _ = compare0' _ (Eq {}) _ _ = pure ()
local {mode := Equal} $ compare0 body1 body2
compareN' s@(Lam {}) t _ _ = clashT s t
compareN' (Sig _ fst1 snd1) (Sig _ fst2 snd2) _ _ = do compare0' ctx ty@(E _) s t = do
compare0 fst1 fst2 -- a neutral type can only be inhabited by neutral values
compare0 snd1 snd2 -- e.g. an abstract value in an abstract type, bound variables, …
compareN' s@(Sig {}) t _ _ = clashT s t E e <- pure s | _ => throwError $ WrongType ty s t
E f <- pure t | _ => throwError $ WrongType ty s t
ignore $ compare0 ctx e f
compareN' (Pair fst1 snd1) (Pair fst2 snd2) _ _ = export covering
local {mode := Equal} $ do compareType : TContext q 0 n -> (s, t : Term q 0 n) -> m ()
compare0 fst1 fst2 compareType ctx s t = do
compare0 snd1 snd2 let Element s ns = whnfD defs s
compareN' s@(Pair {}) t _ _ = clashT s t Element t nt = whnfD defs t
sok <- ensureType s
tok <- ensureType t
compareType' ctx s t
compareN' (Eq _ ty1 l1 r1) (Eq _ ty2 l2 r2) _ _ = do private covering
compare0 ty1 ty2 compareType' : TContext q 0 n -> (s, t : Term q 0 n) ->
local {mode := Equal} $ do (0 ns : NotRedex defs s) => (0 ts : So (isType s)) =>
compare0 l1 l2 (0 nt : NotRedex defs t) => (0 tt : So (isType t)) =>
compare0 r1 r2 m ()
compareN' s@(Eq {}) t _ _ = clashT s t compareType' ctx s t = do
let err : m () = clashT (TYPE UAny) s t
case s of
TYPE k => do
TYPE l <- pure t | _ => err
expectModeU !mode k l
compareN' (DLam _ body1) (DLam _ body2) _ _ = Pi {qty = sQty, arg = sArg, res = sRes, _} => do
local {mode := Equal} $ do Pi {qty = tQty, arg = tArg, res = tRes, _} <- pure t | _ => err
compare0 body1 body2 expectEqualQ sQty tQty
compareN' s@(DLam {}) t _ _ = clashT s t compareType ctx tArg sArg -- contra
-- [todo] is using sArg also ok for subtyping?
compareType (ctx :< sArg) sRes.term tRes.term
compareN' (E e) (E f) ne nf = compareN' e f (noOr2 ne) (noOr2 nf) Sig {fst = sFst, snd = sSnd, _} => do
compareN' s@(E e) t _ _ = clashT s t Sig {fst = tFst, snd = tSnd, _} <- pure t | _ => err
compareType ctx sFst tFst
compareType (ctx :< sFst) sSnd.term tSnd.term
Eq {ty = sTy, l = sl, r = sr, _} => do
Eq {ty = tTy, l = tl, r = tr, _} <- pure t | _ => err
compareType ctx sTy.zero tTy.zero
compareType ctx sTy.one tTy.one
local {mode := Equal} $ do
compare0 ctx sTy.zero sl tl
compare0 ctx sTy.one sr tr
E e => do
E f <- pure t | _ => err
-- no fanciness needed here cos anything other than a neutral
-- has been inlined by whnfD
ignore $ compare0 ctx e f
namespace Elim namespace Elim
export covering export covering %inline
compareN' : CanEqual q m => Eq q => compare0 : TContext q 0 n -> (e, f : Elim q 0 n) -> m (Term q 0 n)
compare0 ctx e f =
let Element e ne = whnfD defs e
Element f nf = whnfD defs f
in
compare0' ctx e f
private
isSubSing : Term {} -> Bool
isSubSing (TYPE _) = False
isSubSing (Pi {res, _}) = isSubSing res.term
isSubSing (Lam {}) = False
isSubSing (Sig {fst, snd, _}) = isSubSing fst && isSubSing snd.term
isSubSing (Pair {}) = False
isSubSing (Eq {}) = True
isSubSing (DLam {}) = False
isSubSing (E e) = False
isSubSing (CloT tm th) = False
isSubSing (DCloT tm th) = False
private covering
compare0' : TContext q 0 n ->
(e, f : Elim q 0 n) -> (e, f : Elim q 0 n) ->
(0 _ : NotRedex defs e) -> (0 _ : NotRedex defs f) -> (0 ne : NotRedex defs e) => (0 nf : NotRedex defs f) =>
m () m (Term q 0 n)
compare0' _ e@(F x) f@(F y) = do
d <- lookupFree' defs x
let ty = d.type
-- [fixme] there is a better way to do this for sure
unless (isSubSing ty.get0 || x == y) $ clashE e f
pure ty.get
compare0' _ e@(F _) f = clashE e f
compareN' e@(F x) f@(F y) _ _ = compare0' ctx e@(B i) f@(B j) = do
unless (x == y) $ clashE e f let ty = ctx !! i
compareN' e@(F _) f _ _ = clashE e f -- [fixme] there is a better way to do this for sure
unless (isSubSing ty || i == j) $ clashE e f
pure ty
compare0' _ e@(B _) f = clashE e f
compareN' e@(B i) f@(B j) _ _ = compare0' ctx (e :@ s) (f :@ t) = local {mode := Equal} $ do
unless (i == j) $ clashE e f Pi {arg, res, _} <- compare0 ctx e f
compareN' e@(B _) f _ _ = clashE e f | ty => throwError $ ExpectedPi ty
compare0 ctx arg s t
pure $ sub1 res (s :# arg)
compare0' _ e@(_ :@ _) f = clashE e f
-- [todo] tracking variance of functions? maybe??? compare0' ctx (CasePair epi e _ eret _ _ ebody)
-- probably not (CasePair fpi f _ fret _ _ fbody) =
compareN' (fun1 :@ arg1) (fun2 :@ arg2) _ _ =
local {mode := Equal} $ do local {mode := Equal} $ do
compare0 fun1 fun2 ty@(Sig {fst, snd, _}) <- compare0 ctx e f
compare0 arg1 arg2 | ty => throwError $ ExpectedSig ty
compareN' e@(_ :@ _) f _ _ = clashE e f unless (epi == fpi) $ throwError $ ClashQ epi fpi
compareType (ctx :< ty) eret.term fret.term
compare0 (ctx :< fst :< snd.term) (substCasePairRet ty eret)
ebody.term fbody.term
pure $ sub1 eret e
compare0' _ e@(CasePair {}) f = clashE e f
compareN' (CasePair pi1 pair1 _ ret1 _ _ body1) compare0' ctx (e :% p) (f :% q) = local {mode := Equal} $ do
(CasePair pi2 pair2 _ ret2 _ _ body2) _ _ = Eq {ty, _} <- compare0 ctx e f
local {mode := Equal} $ do | ty => throwError $ ExpectedEq ty
expectEqualQ pi1 pi2 unless (p == q) $ throwError $ ClashD p q
compare0 pair1 pair2 pure $ dsub1 ty p
compare0 ret1 ret2 compare0' _ e@(_ :% _) f = clashE e f
compare0 body1 body2
compareN' e@(CasePair {}) f _ _ = clashE e f
-- retain the mode unlike above because dimensions can't do compare0' ctx (s :# a) (t :# b) = do
-- anything that would mess up variance compareType ctx a b
compareN' (fun1 :% dim1) (fun2 :% dim2) _ _ = do compare0 ctx a s t
compare0 fun1 fun2 pure b
expectEqualD dim1 dim2 compare0' _ e@(_ :# _) f = clashE e f
compareN' e@(_ :% _) f _ _ = clashE e f
-- using the same mode for the type allows, e.g.
-- A : ★₁ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B
-- which, since A : ★₁ implies A : ★₃, should be fine
compareN' (tm1 :# ty1) (tm2 :# ty2) _ _ = do
compare0 tm1 tm2
compare0 ty1 ty2
compareN' e@(_ :# _) f _ _ = clashE e f
parameters {auto _ : (HasDefs' q _ m, HasErr q m, Eq q)}
(eq : DimEq d) (ctx : TContext q d n)
parameters (mode : EqMode)
namespace Term namespace Term
export covering %inline export covering
compareN : CanEqual q m => Eq q => compare : (ty, s, t : Term q d n) -> m ()
NonRedexTerm q 0 n defs -> NonRedexTerm q 0 n defs -> m () compare ty s t = do
compareN s t = compareN' s.fst t.fst s.snd t.snd defs <- ask
runReaderT {m} (MakeEnv {mode}) $
for_ (splits eq) $ \th =>
compare0 defs (map (/// th) ctx) (ty /// th) (s /// th) (t /// th)
export covering %inline export covering
compare : CanEqual q m => Eq q => compareType : (s, t : Term q d n) -> m ()
DimEq d -> Term q d n -> Term q d n -> m () compareType s t = do
compare eqs s t = defs <- ask
for_ (splits eqs) $ \th => compare0 (s /// th) (t /// th) runReaderT {m} (MakeEnv {mode}) $
for_ (splits eq) $ \th =>
export covering %inline compareType defs (map (/// th) ctx) (s /// th) (t /// th)
compare0 : CanEqual q m => Eq q => Term q 0 n -> Term q 0 n -> m ()
compare0 s t = compareN (whnf defs s) (whnf defs t)
namespace Elim namespace Elim
covering %inline -- can't return the type since it might be different in each dsubst ☹
compareN : CanEqual q m => Eq q =>
NonRedexElim q 0 n defs -> NonRedexElim q 0 n defs -> m ()
compareN e f = compareN' e.fst f.fst e.snd f.snd
export covering %inline export covering %inline
compare : CanEqual q m => Eq q => compare : (e, f : Elim q d n) -> m ()
DimEq d -> Elim q d n -> Elim q d n -> m () compare e f = do
compare eqs e f = defs <- ask
for_ (splits eqs) $ \th => compare0 (e /// th) (f /// th) runReaderT {m} (MakeEnv {mode}) $
for_ (splits eq) $ \th =>
export covering %inline ignore $ compare0 defs (map (/// th) ctx) (e /// th) (f /// th)
compare0 : CanEqual q m => Eq q => Elim q 0 n -> Elim q 0 n -> m ()
compare0 e f = compareN (whnf defs e) (whnf defs f)
namespace ScopeTermN
export covering %inline
compare0 : {s : Nat} -> CanEqual q m => Eq q =>
ScopeTermN s q 0 n -> ScopeTermN s q 0 n -> m ()
compare0 (TUnused body0) (TUnused body1) = compare0 body0 body1
compare0 body0 body1 = compare0 body0.term body1.term
-- [todo] extend to multi-var scopes
namespace DScopeTerm
export covering %inline
compare0 : CanEqual q m => Eq q =>
DScopeTerm q 0 n -> DScopeTerm q 0 n -> m ()
compare0 (DUnused body0) (DUnused body1) = compare0 body0 body1
compare0 body0 body1 = do
compare0 body0.zero body1.zero
compare0 body0.one body1.one
namespace Term namespace Term
export covering %inline export covering %inline
equal : HasErr q m => Eq q => equal, sub : (ty, s, t : Term q d n) -> m ()
DimEq d -> Term q d n -> Term q d n -> m () equal = compare Equal
equal eqs s t {m} = runReaderT {m} (MakeEnv Equal) $ compare eqs s t sub = compare Sub
export covering %inline export covering %inline
sub : HasErr q m => HasDefs' q _ m => Eq q => equalType, subtype : (s, t : Term q d n) -> m ()
DimEq d -> Term q d n -> Term q d n -> m () equalType = compareType Equal
sub eqs s t {m} = runReaderT {m} (MakeEnv Sub) $ compare eqs s t subtype = compareType Sub
namespace Elim namespace Elim
export covering %inline export covering %inline
equal : HasErr q m => Eq q => equal, sub : (e, f : Elim q d n) -> m ()
DimEq d -> Elim q d n -> Elim q d n -> m () equal = compare Equal
equal eqs e f {m} = runReaderT {m} (MakeEnv Equal) $ compare eqs e f sub = compare Sub
export covering %inline
sub : HasErr q m => HasDefs' q _ m => Eq q =>
DimEq d -> Elim q d n -> Elim q d n -> m ()
sub eqs e f {m} = runReaderT {m} (MakeEnv Sub) $ compare eqs e f

View file

@ -125,15 +125,6 @@ parameters (th : DSubst dfrom dto) (ph : TSubst q dto from to)
pushSubstsWith' e = (pushSubstsWith th ph e).fst pushSubstsWith' e = (pushSubstsWith th ph e).fst
public export %inline
weakT : Term q d n -> Term q d (S n)
weakT t = t //. shift 1
public export %inline
weakE : Elim q d n -> Elim q d (S n)
weakE t = t //. shift 1
public export 0 public export 0
Lookup : TermLike Lookup : TermLike
Lookup q d n = Name -> Maybe $ Elim q d n Lookup q d n = Name -> Maybe $ Elim q d n

View file

@ -184,17 +184,35 @@ comp : DSubst dfrom dto -> TSubst q dfrom from mid -> TSubst q dto mid to ->
comp th ps ph = map (/// th) ps . ph 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
namespace ScopeTermN namespace ScopeTermN
export %inline export %inline
(.term) : {s : Nat} -> ScopeTermN s q d n -> Term q d (s + n) (.term) : {s : Nat} -> ScopeTermN s q d n -> Term q d (s + n)
(TUsed body).term = body (TUsed body).term = body
(TUnused body).term = body //. shift s (TUnused body).term = weakT body {by = s}
namespace DScopeTermN namespace DScopeTermN
export %inline export %inline
(.term) : {s : Nat} -> DScopeTermN s q d n -> Term q (s + d) n (.term) : {s : Nat} -> DScopeTermN s q d n -> Term q (s + d) n
(DUsed body).term = body (DUsed body).term = body
(DUnused body).term = body /// shift s (DUnused body).term = dweakT body {by = s}
export %inline export %inline

View file

@ -22,7 +22,7 @@ CanTC q = CanTC' q IsGlobal
private covering %inline private covering %inline
expectTYPE : CanTC' q _ m => Term q d n -> m Universe expectTYPE : CanTC' q _ m => Term q d n -> m Universe
expectTYPE s = expectTYPE s =
case whnf !ask s of case whnfD !ask s of
Element (TYPE l) _ => pure l Element (TYPE l) _ => pure l
_ => throwError $ ExpectedTYPE s _ => throwError $ ExpectedTYPE s
@ -30,7 +30,7 @@ private covering %inline
expectPi : CanTC' q _ m => Term q d n -> expectPi : CanTC' q _ m => Term q d n ->
m (q, Term q d n, ScopeTerm q d n) m (q, Term q d n, ScopeTerm q d n)
expectPi ty = expectPi ty =
case whnf !ask ty of case whnfD !ask ty of
Element (Pi qty _ arg res) _ => pure (qty, arg, res) Element (Pi qty _ arg res) _ => pure (qty, arg, res)
_ => throwError $ ExpectedPi ty _ => throwError $ ExpectedPi ty
@ -38,7 +38,7 @@ private covering %inline
expectSig : CanTC' q _ m => Term q d n -> expectSig : CanTC' q _ m => Term q d n ->
m (Term q d n, ScopeTerm q d n) m (Term q d n, ScopeTerm q d n)
expectSig ty = expectSig ty =
case whnf !ask ty of case whnfD !ask ty of
Element (Sig _ arg res) _ => pure (arg, res) Element (Sig _ arg res) _ => pure (arg, res)
_ => throwError $ ExpectedSig ty _ => throwError $ ExpectedSig ty
@ -46,7 +46,7 @@ private covering %inline
expectEq : CanTC' q _ m => Term q d n -> expectEq : CanTC' q _ m => Term q d n ->
m (DScopeTerm q d n, Term q d n, Term q d n) m (DScopeTerm q d n, Term q d n, Term q d n)
expectEq ty = expectEq ty =
case whnf !ask ty of case whnfD !ask ty of
Element (Eq _ ty l r) _ => pure (ty, l, r) Element (Eq _ ty l r) _ => pure (ty, l, r)
_ => throwError $ ExpectedEq ty _ => throwError $ ExpectedEq ty
@ -80,11 +80,8 @@ lookupBound pi (VS i) ctx =
weakI $ lookupBound pi i (tail ctx) weakI $ lookupBound pi i (tail ctx)
private private
lookupFree : IsQty q => CanTC q m => Name -> m (Definition q) lookupFree : CanTC' q g m => Name -> m (Definition' q g)
lookupFree x = lookupFree x = lookupFree' !ask x
case lookup x !ask of
Just d => pure d
Nothing => throwError $ NotInScope x
private %inline private %inline
subjMult : IsQty q => (sg : SQty q) -> q -> SQty q subjMult : IsQty q => (sg : SQty q) -> q -> SQty q
@ -202,9 +199,9 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
qout <- check (extendDim ctx) sg body.term ty.term qout <- check (extendDim ctx) sg body.term ty.term
let eqs = makeDimEq ctx.dctx let eqs = makeDimEq ctx.dctx
-- if Ψ ⊢ t0 = l -- if Ψ ⊢ t0 = l
equal !ask eqs body.zero l equal eqs ctx.tctx ty.zero body.zero l
-- if Ψ ⊢ t1 = r -- if Ψ ⊢ t1 = r
equal !ask eqs body.one r equal eqs ctx.tctx ty.one body.one r
-- then Ψ | Γ ⊢ (λᴰi ⇒ t) · σ ⇐ Eq [i ⇒ A] l r ⊳ Σ -- then Ψ | Γ ⊢ (λᴰi ⇒ t) · σ ⇐ Eq [i ⇒ A] l r ⊳ Σ
pure qout pure qout
@ -212,7 +209,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- if Ψ | Γ ⊢ e · σ ⇒ A' ⊳ Σ -- if Ψ | Γ ⊢ e · σ ⇒ A' ⊳ Σ
infres <- infer ctx sg e infres <- infer ctx sg e
-- if Ψ ⊢ A' <: A -- if Ψ ⊢ A' <: A
sub !ask (makeDimEq ctx.dctx) infres.type ty subtype (makeDimEq ctx.dctx) ctx.tctx infres.type ty
-- then Ψ | Γ ⊢ e · σ ⇐ A ⊳ Σ -- then Ψ | Γ ⊢ e · σ ⇐ A ⊳ Σ
pure infres.qout pure infres.qout
@ -258,8 +255,7 @@ parameters {auto _ : IsQty q} {auto _ : CanTC q m}
-- ⊳ Σ₂, x · π₁, y · π₂ -- ⊳ Σ₂, x · π₁, y · π₂
-- if π₁, π₂ ≤ π -- if π₁, π₂ ≤ π
let bodyctx = extendTyN [< (tfst, pi), (tsnd.term, pi)] ctx let bodyctx = extendTyN [< (tfst, pi), (tsnd.term, pi)] ctx
retarg = Pair (BVT 0) (BVT 1) :# (pairres.type // fromNat 2) bodyty = substCasePairRet pairres.type ret
bodyty = ret.term // (retarg ::: shift 2)
bodyout <- check bodyctx sg body.term bodyty >>= popQs [< pi, pi] bodyout <- check bodyctx sg body.term bodyty >>= popQs [< pi, pi]
-- then Ψ | Γ ⊢ σ case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂ -- then Ψ | Γ ⊢ σ case ⋯ ⇒ ret[pair] ⊳ πΣ₁ + Σ₂
pure $ InfRes { pure $ InfRes {

View file

@ -121,12 +121,17 @@ data Error q
| ExpectedEq (Term q d n) | ExpectedEq (Term q d n)
| BadUniverse Universe Universe | BadUniverse Universe Universe
| ClashT EqMode (Term q d n) (Term q d n) -- first arg of ClashT is the type
| ClashT EqMode (Term q d n) (Term q d n) (Term q d n)
| ClashE EqMode (Elim q d n) (Elim q d n)
| ClashU EqMode Universe Universe | ClashU EqMode Universe Universe
| ClashQ q q | ClashQ q q
| ClashD (Dim d) (Dim d) | ClashD (Dim d) (Dim d)
| NotInScope Name | NotInScope Name
| NotType (Term q d n)
| WrongType (Term q d n) (Term q d n) (Term q d n)
public export public export
0 HasErr : Type -> (Type -> Type) -> Type 0 HasErr : Type -> (Type -> Type) -> Type
HasErr q = MonadError (Error q) HasErr q = MonadError (Error q)
@ -157,3 +162,18 @@ parameters {auto _ : HasErr q m}
export %inline export %inline
expectEqualD : Dim d -> Dim d -> m () expectEqualD : Dim d -> Dim d -> m ()
expectEqualD = expect ClashD (==) expectEqualD = expect ClashD (==)
export
lookupFree' : HasErr q m => Definitions' q g -> Name -> m (Definition' q g)
lookupFree' defs x =
case lookup x defs of
Just d => pure d
Nothing => throwError $ NotInScope x
export
substCasePairRet : Term q d n -> ScopeTerm q d n -> Term q d (2 + n)
substCasePairRet dty retty =
let arg = Pair (BVT 0) (BVT 1) :# (dty // fromNat 2) in
retty.term // (arg ::: shift 2)

View file

@ -26,11 +26,17 @@ ToInfo (Error Three) where
[("type", "BadUniverse"), [("type", "BadUniverse"),
("low", show k), ("low", show k),
("high", show l)] ("high", show l)]
toInfo (ClashT mode s t) = toInfo (ClashT mode ty s t) =
[("type", "ClashT"), [("type", "ClashT"),
("mode", show mode), ("mode", show mode),
("ty", prettyStr True ty),
("left", prettyStr True s), ("left", prettyStr True s),
("right", prettyStr True t)] ("right", prettyStr True t)]
toInfo (ClashE mode e f) =
[("type", "ClashE"),
("mode", show mode),
("left", prettyStr True e),
("right", prettyStr True f)]
toInfo (ClashU mode k l) = toInfo (ClashU mode k l) =
[("type", "ClashU"), [("type", "ClashU"),
("mode", show mode), ("mode", show mode),
@ -44,13 +50,29 @@ ToInfo (Error Three) where
[("type", "ClashD"), [("type", "ClashD"),
("left", prettyStr True p), ("left", prettyStr True p),
("right", prettyStr True q)] ("right", prettyStr True q)]
toInfo (NotType ty) =
[("type", "NotType"),
("actual", prettyStr True ty)]
toInfo (WrongType ty s t) =
[("type", "WrongType"),
("ty", prettyStr True ty),
("left", prettyStr True s),
("right", prettyStr True t)]
0 M : Type -> Type 0 M : Type -> Type
M = ReaderT (Definitions Three) (Either (Error Three)) M = ReaderT (Definitions Three) (Either (Error Three))
defGlobals : Definitions Three
defGlobals = fromList
[("A", mkAbstract Zero $ TYPE 0),
("B", mkAbstract Zero $ TYPE 0),
("a", mkAbstract Any $ FT "A"),
("b", mkAbstract Any $ FT "B"),
("f", mkAbstract Any $ Arr One (FT "A") (FT "A"))]
parameters (label : String) (act : Lazy (M ())) parameters (label : String) (act : Lazy (M ()))
{default empty globals : Definitions Three} {default defGlobals globals : Definitions Three}
testEq : Test testEq : Test
testEq = test label $ runReaderT globals act testEq = test label $ runReaderT globals act
@ -60,17 +82,18 @@ parameters (label : String) (act : Lazy (M ()))
parameters {default 0 d, n : Nat} parameters {default 0 d, n : Nat}
{default new eqs : DimEq d} {default new eqs : DimEq d}
subT : Term Three d n -> Term Three d n -> M () (ctx : TContext Three d n)
subT s t = Term.sub !ask eqs s t subT : Term Three d n -> Term Three d n -> Term Three d n -> M ()
subT ty s t = Term.sub eqs ctx ty s t
equalT : Term Three d n -> Term Three d n -> M () equalT : Term Three d n -> Term Three d n -> Term Three d n -> M ()
equalT s t = Term.equal !ask eqs s t equalT ty s t = Term.equal eqs ctx ty s t
subE : Elim Three d n -> Elim Three d n -> M () subE : Elim Three d n -> Elim Three d n -> M ()
subE e f = Elim.sub !ask eqs e f subE e f = Elim.sub eqs ctx e f
equalE : Elim Three d n -> Elim Three d n -> M () equalE : Elim Three d n -> Elim Three d n -> M ()
equalE e f = Elim.equal !ask eqs e f equalE e f = Elim.equal eqs ctx e f
export export
@ -80,17 +103,17 @@ tests = "equality & subtyping" :- [
"universes" :- [ "universes" :- [
testEq "★₀ ≡ ★₀" $ testEq "★₀ ≡ ★₀" $
equalT (TYPE 0) (TYPE 0), equalT [<] (TYPE 1) (TYPE 0) (TYPE 0),
testNeq "★₀ ≢ ★₁" $ testNeq "★₀ ≢ ★₁" $
equalT (TYPE 0) (TYPE 1), equalT [<] (TYPE 2) (TYPE 0) (TYPE 1),
testNeq "★₁ ≢ ★₀" $ testNeq "★₁ ≢ ★₀" $
equalT (TYPE 1) (TYPE 0), equalT [<] (TYPE 2) (TYPE 1) (TYPE 0),
testEq "★₀ <: ★₀" $ testEq "★₀ <: ★₀" $
subT (TYPE 0) (TYPE 0), subT [<] (TYPE 1) (TYPE 0) (TYPE 0),
testEq "★₀ <: ★₁" $ testEq "★₀ <: ★₁" $
subT (TYPE 0) (TYPE 1), subT [<] (TYPE 2) (TYPE 0) (TYPE 1),
testNeq "★₁ ≮: ★₀" $ testNeq "★₁ ≮: ★₀" $
subT (TYPE 1) (TYPE 0) subT [<] (TYPE 2) (TYPE 1) (TYPE 0)
], ],
"pi" :- [ "pi" :- [
@ -98,79 +121,91 @@ tests = "equality & subtyping" :- [
note #""AB" for (0 _ : A) → B"#, note #""AB" for (0 _ : A) → B"#,
testEq "A ⊸ B ≡ A ⊸ B" $ testEq "A ⊸ B ≡ A ⊸ B" $
let tm = Arr One (FT "A") (FT "B") in let tm = Arr One (FT "A") (FT "B") in
equalT tm tm, equalT [<] (TYPE 0) tm tm,
testNeq "A ⇾ B ≢ A ⊸ B" $ testNeq "A ⇾ B ≢ A ⊸ B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = Arr Zero (FT "A") (FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = Arr One (FT "A") (FT "B") in
equalT tm1 tm2, equalT [<] (TYPE 0) tm1 tm2,
testEq "0=1 ⊢ A ⇾ B ≢ A ⊸ B" $ testEq "0=1 ⊢ A ⇾ B ≢ A ⊸ B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = Arr Zero (FT "A") (FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = Arr One (FT "A") (FT "B") in
equalT tm1 tm2 {eqs = ZeroIsOne}, equalT [<] (TYPE 0) tm1 tm2 {eqs = ZeroIsOne},
testEq "A ⊸ B <: A ⊸ B" $ testEq "A ⊸ B <: A ⊸ B" $
let tm = Arr One (FT "A") (FT "B") in let tm = Arr One (FT "A") (FT "B") in
subT tm tm, subT [<] (TYPE 0) tm tm,
testNeq "A ⇾ B ≮: A ⊸ B" $ testNeq "A ⇾ B ≮: A ⊸ B" $
let tm1 = Arr Zero (FT "A") (FT "B") let tm1 = Arr Zero (FT "A") (FT "B")
tm2 = Arr One (FT "A") (FT "B") in tm2 = Arr One (FT "A") (FT "B") in
subT tm1 tm2, subT [<] (TYPE 0) tm1 tm2,
testEq "★₀ ⇾ ★₀ ≡ ★₀ ⇾ ★₀" $ testEq "★₀ ⇾ ★₀ ≡ ★₀ ⇾ ★₀" $
let tm = Arr Zero (TYPE 0) (TYPE 0) in let tm = Arr Zero (TYPE 0) (TYPE 0) in
equalT tm tm, equalT [<] (TYPE 1) tm tm,
testEq "★₀ ⇾ ★₀ <: ★₀ ⇾ ★₀" $ testEq "★₀ ⇾ ★₀ <: ★₀ ⇾ ★₀" $
let tm = Arr Zero (TYPE 0) (TYPE 0) in let tm = Arr Zero (TYPE 0) (TYPE 0) in
subT tm tm, subT [<] (TYPE 1) tm tm,
testNeq "★₁ ⊸ ★₀ ≢ ★₀ ⇾ ★₀" $ testNeq "★₁ ⊸ ★₀ ≢ ★₀ ⇾ ★₀" $
let tm1 = Arr Zero (TYPE 1) (TYPE 0) let tm1 = Arr Zero (TYPE 1) (TYPE 0)
tm2 = Arr Zero (TYPE 0) (TYPE 0) in tm2 = Arr Zero (TYPE 0) (TYPE 0) in
equalT tm1 tm2, equalT [<] (TYPE 2) tm1 tm2,
testEq "★₁ ⊸ ★₀ <: ★₀ ⊸ ★₀" $ testEq "★₁ ⊸ ★₀ <: ★₀ ⊸ ★₀" $
let tm1 = Arr One (TYPE 1) (TYPE 0) let tm1 = Arr One (TYPE 1) (TYPE 0)
tm2 = Arr One (TYPE 0) (TYPE 0) in tm2 = Arr One (TYPE 0) (TYPE 0) in
subT tm1 tm2, subT [<] (TYPE 2) tm1 tm2,
testNeq "★₀ ⊸ ★₀ ≢ ★₀ ⇾ ★₁" $ testNeq "★₀ ⊸ ★₀ ≢ ★₀ ⇾ ★₁" $
let tm1 = Arr Zero (TYPE 0) (TYPE 0) let tm1 = Arr Zero (TYPE 0) (TYPE 0)
tm2 = Arr Zero (TYPE 0) (TYPE 1) in tm2 = Arr Zero (TYPE 0) (TYPE 1) in
equalT tm1 tm2, equalT [<] (TYPE 2) tm1 tm2,
testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $ testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $
let tm1 = Arr One (TYPE 0) (TYPE 0) let tm1 = Arr One (TYPE 0) (TYPE 0)
tm2 = Arr One (TYPE 0) (TYPE 1) in tm2 = Arr One (TYPE 0) (TYPE 1) in
subT tm1 tm2, subT [<] (TYPE 2) tm1 tm2,
testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $ testEq "★₀ ⊸ ★₀ <: ★₀ ⊸ ★₁" $
let tm1 = Arr One (TYPE 0) (TYPE 0) let tm1 = Arr One (TYPE 0) (TYPE 0)
tm2 = Arr One (TYPE 0) (TYPE 1) in tm2 = Arr One (TYPE 0) (TYPE 1) in
subT tm1 tm2 subT [<] (TYPE 2) tm1 tm2
], ],
"lambda" :- [ "lambda" :- [
testEq "λ x ⇒ [x] ≡ λ x ⇒ [x]" $ testEq "λ x ⇒ [x] ≡ λ x ⇒ [x]" $
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "x" $ TUsed $ BVT 0), equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "x" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] <: λ x ⇒ [x]" $ testEq "λ x ⇒ [x] <: λ x ⇒ [x]" $
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "x" $ TUsed $ BVT 0), subT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "x" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] ≡ λ y ⇒ [y]" $ testEq "λ x ⇒ [x] ≡ λ y ⇒ [y]" $
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "y" $ TUsed $ BVT 0), equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "y" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] <: λ y ⇒ [y]" $ testEq "λ x ⇒ [x] <: λ y ⇒ [y]" $
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "y" $ TUsed $ BVT 0), equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "y" $ TUsed $ BVT 0),
testNeq "λ x y ⇒ [x] ≢ λ x y ⇒ [y]" $ testNeq "λ x y ⇒ [x] ≢ λ x y ⇒ [y]" $
equalT (Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 1) equalT [<] (Arr One (FT "A") $ Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 0), (Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 1)
(Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 0),
testEq "λ x ⇒ [a] ≡ λ x ⇒ [a] (TUsed vs TUnused)" $ testEq "λ x ⇒ [a] ≡ λ x ⇒ [a] (TUsed vs TUnused)" $
equalT (Lam "x" $ TUsed $ FT "a") equalT [<] (Arr Zero (FT "B") (FT "A"))
(Lam "x" $ TUnused $ FT "a"), (Lam "x" $ TUsed $ FT "a")
(Lam "x" $ TUnused $ FT "a"),
skipWith "(no η yet)" $ skipWith "(no η yet)" $
testEq "λ x ⇒ [f [x]] ≡ [f] (η)" $ testEq "λ x ⇒ [f [x]] ≡ [f] (η)" $
equalT (Lam "x" $ TUsed $ E $ F "f" :@ BVT 0) equalT [<] (Arr One (FT "A") (FT "A"))
(FT "f") (Lam "x" $ TUsed $ E $ F "f" :@ BVT 0)
(FT "f")
], ],
"eq type" :- [ "eq type" :- [
testEq "(★₀ = ★₀ : ★₁) ≡ (★₀ = ★₀ : ★₁)" $ testEq "(★₀ = ★₀ : ★₁) ≡ (★₀ = ★₀ : ★₁)" $
let tm = Eq0 (TYPE 1) (TYPE 0) (TYPE 0) in let tm = Eq0 (TYPE 1) (TYPE 0) (TYPE 0) in
equalT tm tm, equalT [<] (TYPE 2) tm tm,
testEq "A ≔ ★₁ ⊢ (★₀ = ★₀ : ★₁) ≡ (★₀ = ★₀ : A)" testEq "A ≔ ★₁ ⊢ (★₀ = ★₀ : ★₁) ≡ (★₀ = ★₀ : A)"
{globals = fromList [("A", mkDef zero (TYPE 2) (TYPE 1))]} $ {globals = fromList [("A", mkDef zero (TYPE 2) (TYPE 1))]} $
equalT (Eq0 (TYPE 1) (TYPE 0) (TYPE 0)) equalT [<] (TYPE 2)
(Eq0 (FT "A") (TYPE 0) (TYPE 0)) (Eq0 (TYPE 1) (TYPE 0) (TYPE 0))
(Eq0 (FT "A") (TYPE 0) (TYPE 0))
], ],
todo "dim lambda", todo "dim lambda",
@ -178,17 +213,25 @@ tests = "equality & subtyping" :- [
"term closure" :- [ "term closure" :- [
note "𝑖, 𝑗 for bound variables pointing outside of the current expr", note "𝑖, 𝑗 for bound variables pointing outside of the current expr",
testEq "[𝑖]{} ≡ [𝑖]" $ testEq "[𝑖]{} ≡ [𝑖]" $
equalT (CloT (BVT 0) id) (BVT 0) {n = 1}, equalT [< FT "A"] (FT "A") {n = 1}
(CloT (BVT 0) id)
(BVT 0),
testEq "[𝑖]{a/𝑖} ≡ [a]" $ testEq "[𝑖]{a/𝑖} ≡ [a]" $
equalT (CloT (BVT 0) (F "a" ::: id)) (FT "a"), equalT [<] (FT "A")
(CloT (BVT 0) (F "a" ::: id))
(FT "a"),
testEq "[𝑖]{a/𝑖,b/𝑗} ≡ [a]" $ testEq "[𝑖]{a/𝑖,b/𝑗} ≡ [a]" $
equalT (CloT (BVT 0) (F "a" ::: F "b" ::: id)) (FT "a"), equalT [<] (FT "A")
(CloT (BVT 0) (F "a" ::: F "b" ::: id))
(FT "a"),
testEq "(λy. [𝑖]){y/y, a/𝑖} ≡ λy. [a] (TUnused)" $ testEq "(λy. [𝑖]){y/y, a/𝑖} ≡ λy. [a] (TUnused)" $
equalT (CloT (Lam "y" $ TUnused $ BVT 0) (F "a" ::: id)) equalT [<] (Arr Zero (FT "B") (FT "A"))
(Lam "y" $ TUnused $ FT "a"), (CloT (Lam "y" $ TUnused $ BVT 0) (F "a" ::: id))
(Lam "y" $ TUnused $ FT "a"),
testEq "(λy. [𝑖]){y/y, a/𝑖} ≡ λy. [a] (TUsed)" $ testEq "(λy. [𝑖]){y/y, a/𝑖} ≡ λy. [a] (TUsed)" $
equalT (CloT (Lam "y" $ TUsed $ BVT 1) (F "a" ::: id)) equalT [<] (Arr Zero (FT "B") (FT "A"))
(Lam "y" $ TUsed $ FT "a") (CloT (Lam "y" $ TUsed $ BVT 1) (F "a" ::: id))
(Lam "y" $ TUsed $ FT "a")
], ],
todo "term d-closure", todo "term d-closure",
@ -202,73 +245,74 @@ tests = "equality & subtyping" :- [
("B", mkDef Any (TYPE (U 1)) (FT "A"))] ("B", mkDef Any (TYPE (U 1)) (FT "A"))]
in [ in [
testEq "A ≡ A" $ testEq "A ≡ A" $
equalE (F "A") (F "A"), equalE [<] (F "A") (F "A"),
testNeq "A ≢ B" $ testNeq "A ≢ B" $
equalE (F "A") (F "B"), equalE [<] (F "A") (F "B"),
testEq "0=1 ⊢ A ≡ B" $ testEq "0=1 ⊢ A ≡ B" $
equalE {eqs = ZeroIsOne} (F "A") (F "B"), equalE {eqs = ZeroIsOne} [<] (F "A") (F "B"),
testEq "A : ★₁ ≔ ★₀ ⊢ A ≡ (★₀ ∷ ★₁)" {globals = au_bu} $ testEq "A : ★₁ ≔ ★₀ ⊢ A ≡ (★₀ ∷ ★₁)" {globals = au_bu} $
equalE (F "A") (TYPE 0 :# TYPE 1), equalE [<] (F "A") (TYPE 0 :# TYPE 1),
testEq "A ≔ ★₀, B ≔ ★₀ ⊢ A ≡ B" {globals = au_bu} $ testEq "A ≔ ★₀, B ≔ ★₀ ⊢ A ≡ B" {globals = au_bu} $
equalE (F "A") (F "B"), equalE [<] (F "A") (F "B"),
testEq "A ≔ ★₀, B ≔ A ⊢ A ≡ B" {globals = au_ba} $ testEq "A ≔ ★₀, B ≔ A ⊢ A ≡ B" {globals = au_ba} $
equalE (F "A") (F "B"), equalE [<] (F "A") (F "B"),
testEq "A <: A" $ testEq "A <: A" $
subE (F "A") (F "A"), subE [<] (F "A") (F "A"),
testNeq "A ≮: B" $ testNeq "A ≮: B" $
subE (F "A") (F "B"), subE [<] (F "A") (F "B"),
testEq "A : ★₃ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B" testEq "A : ★₃ ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B"
{globals = fromList [("A", mkDef Any (TYPE 3) (TYPE 0)), {globals = fromList [("A", mkDef Any (TYPE 3) (TYPE 0)),
("B", mkDef Any (TYPE 3) (TYPE 2))]} $ ("B", mkDef Any (TYPE 3) (TYPE 2))]} $
subE (F "A") (F "B"), subE [<] (F "A") (F "B"),
testEq "A : ★₁👈 ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B" testEq "A : ★₁👈 ≔ ★₀, B : ★₃ ≔ ★₂ ⊢ A <: B"
{globals = fromList [("A", mkDef Any (TYPE 1) (TYPE 0)), {globals = fromList [("A", mkDef Any (TYPE 1) (TYPE 0)),
("B", mkDef Any (TYPE 3) (TYPE 2))]} $ ("B", mkDef Any (TYPE 3) (TYPE 2))]} $
subE (F "A") (F "B"), subE [<] (F "A") (F "B"),
testEq "0=1 ⊢ A <: B" $ testEq "0=1 ⊢ A <: B" $
subE (F "A") (F "B") {eqs = ZeroIsOne} subE [<] (F "A") (F "B") {eqs = ZeroIsOne}
], ],
"bound var" :- [ "bound var" :- [
note "𝑖, 𝑗 for distinct bound variables", note "𝑖, 𝑗 for distinct bound variables",
testEq "𝑖𝑖" $ testEq "𝑖𝑖" $
equalE (BV 0) (BV 0) {n = 1}, equalE [< TYPE 0] (BV 0) (BV 0) {n = 1},
testNeq "𝑖𝑗" $ testNeq "𝑖𝑗" $
equalE (BV 0) (BV 1) {n = 2}, equalE [< TYPE 0, TYPE 0] (BV 0) (BV 1) {n = 2},
testEq "0=1 ⊢ 𝑖𝑗" $ testEq "0=1 ⊢ 𝑖𝑗" $
equalE {n = 2, eqs = ZeroIsOne} (BV 0) (BV 1) equalE [< TYPE 0, TYPE 0] (BV 0) (BV 1)
{n = 2, eqs = ZeroIsOne}
], ],
"application" :- [ "application" :- [
testEq "f [a] ≡ f [a]" $ testEq "f [a] ≡ f [a]" $
equalE (F "f" :@ FT "a") (F "f" :@ FT "a"), equalE [<] (F "f" :@ FT "a") (F "f" :@ FT "a"),
testEq "f [a] <: f [a]" $ testEq "f [a] <: f [a]" $
subE (F "f" :@ FT "a") (F "f" :@ FT "a"), subE [<] (F "f" :@ FT "a") (F "f" :@ FT "a"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ ([a ∷ A] ∷ A) (β)" $ testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ ([a ∷ A] ∷ A) (β)" $
equalE equalE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A"))) ((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a") :@ FT "a")
(E (FT "a" :# FT "A") :# FT "A"), (E (FT "a" :# FT "A") :# FT "A"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ a (βυ)" $ testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ a (βυ)" $
equalE equalE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A"))) ((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a") :@ FT "a")
(F "a"), (F "a"),
testEq "(λ g ⇒ [g [x]] ∷ ⋯)) [f] ≡ (λ y ⇒ [f [y]] ∷ ⋯) [x] (β↘↙)" $ testEq "(λ g ⇒ [g [a]] ∷ ⋯)) [f] ≡ (λ y ⇒ [f [y]] ∷ ⋯) [a] (β↘↙)" $
let a = FT "A"; a2a = (Arr One a a) in let a = FT "A"; a2a = (Arr One a a) in
equalE equalE [<]
((Lam "g" (TUsed (E (BV 0 :@ FT "x"))) :# Arr One a2a a) :@ FT "f") ((Lam "g" (TUsed (E (BV 0 :@ FT "a"))) :# Arr One a2a a) :@ FT "f")
((Lam "y" (TUsed (E (F "f" :@ BVT 0))) :# a2a) :@ FT "x"), ((Lam "y" (TUsed (E (F "f" :@ BVT 0))) :# a2a) :@ FT "a"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a <: a" $ testEq "(λ x ⇒ [x] ∷ A ⊸ A) a <: a" $
subE subE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A"))) ((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a") :@ FT "a")
(F "a"), (F "a"),
testEq "f : A ⊸ A ≔ λ x ⇒ [x] ⊢ f [x] ≡ x" testEq "id : A ⊸ A ≔ λ x ⇒ [x] ⊢ id [a] ≡ a"
{globals = fromList {globals = defGlobals `mergeLeft` fromList
[("f", mkDef Any (Arr One (FT "A") (FT "A")) [("id", mkDef Any (Arr One (FT "A") (FT "A"))
(Lam "x" (TUsed (BVT 0))))]} $ (Lam "x" (TUsed (BVT 0))))]} $
equalE (F "f" :@ FT "x") (F "x") equalE [<] (F "id" :@ FT "a") (F "a")
], ],
"dim application" :- "dim application" :-
@ -277,13 +321,16 @@ tests = "equality & subtyping" :- [
in in
[ [
note #""refl [A] x" is an abbreviation for "(λᴰi ⇒ x)(x ≡ x : A)""#, note #""refl [A] x" is an abbreviation for "(λᴰi ⇒ x)(x ≡ x : A)""#,
testEq "refl [A] x ≡ refl [A] x" $ testEq "refl [A] a ≡ refl [A] a" $
equalE (refl (FT "A") (FT "x")) (refl (FT "A") (FT "x")), equalE [<] (refl (FT "A") (FT "a")) (refl (FT "A") (FT "a")),
testEq "p : (a ≡ b : A), q : (a ≡ b : A) ⊢ p ≡ q" testEq "p : (a ≡ b : A), q : (a ≡ b : A) ⊢ p ≡ q"
{globals = {globals =
let def = mkAbstract Zero $ Eq0 (FT "A") (FT "a") (FT "b") in let def = mkAbstract Zero $ Eq0 (FT "A") (FT "a") (FT "b") in
fromList [("p", def), ("q", def)]} $ fromList [("A", mkAbstract Zero $ TYPE 0),
equalE (F "p") (F "q") ("a", mkAbstract Any $ FT "A"),
("b", mkAbstract Any $ FT "A"),
("p", def), ("q", def)]} $
equalE [<] (F "p") (F "q")
], ],
todo "annotation", todo "annotation",
@ -294,9 +341,10 @@ tests = "equality & subtyping" :- [
"clashes" :- [ "clashes" :- [
testNeq "★₀ ≢ ★₀ ⇾ ★₀" $ testNeq "★₀ ≢ ★₀ ⇾ ★₀" $
equalT (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0)), equalT [<] (TYPE 1) (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0)),
testEq "0=1 ⊢ ★₀ ≡ ★₀ ⇾ ★₀" $ testEq "0=1 ⊢ ★₀ ≡ ★₀ ⇾ ★₀" $
equalT (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0)) {eqs = ZeroIsOne}, equalT [<] (TYPE 1) (TYPE 0) (Arr Zero (TYPE 0) (TYPE 0))
{eqs = ZeroIsOne},
todo "others" todo "others"
] ]
] ]

View file

@ -27,16 +27,16 @@ testNoStep whnf label e = test "\{label} (no step)" $
parameters {default empty defs : Definitions Three} {default 0 d, n : Nat} parameters {default empty defs : Definitions Three} {default 0 d, n : Nat}
testWhnfT : String -> Term Three d n -> Term Three d n -> Test testWhnfT : String -> Term Three d n -> Term Three d n -> Test
testWhnfT = testWhnf (whnf defs) testWhnfT = testWhnf (whnfD defs)
testWhnfE : String -> Elim Three d n -> Elim Three d n -> Test testWhnfE : String -> Elim Three d n -> Elim Three d n -> Test
testWhnfE = testWhnf (whnf defs) testWhnfE = testWhnf (whnfD defs)
testNoStepE : String -> Elim Three d n -> Test testNoStepE : String -> Elim Three d n -> Test
testNoStepE = testNoStep (whnf defs) testNoStepE = testNoStep (whnfD defs)
testNoStepT : String -> Term Three d n -> Test testNoStepT : String -> Term Three d n -> Test
testNoStepT = testNoStep (whnf defs) testNoStepT = testNoStep (whnfD defs)
tests = "whnf" :- [ tests = "whnf" :- [