more typed equality, uip, etc

This commit is contained in:
rhiannon morris 2023-02-11 18:15:50 +01:00
parent 7fd7a31635
commit 7d2c3b5a8e
8 changed files with 381 additions and 217 deletions

View file

@ -1,65 +1,10 @@
module Tests.Equal
import Quox.Equal as Lib
import Quox.Pretty
import Quox.Equal
import Quox.Syntax.Qty.Three
import public TypingImpls
import TAP
export
ToInfo (Error Three) where
toInfo (NotInScope x) =
[("type", "NotInScope"),
("name", show x)]
toInfo (ExpectedTYPE t) =
[("type", "ExpectedTYPE"),
("got", prettyStr True t)]
toInfo (ExpectedPi t) =
[("type", "ExpectedPi"),
("got", prettyStr True t)]
toInfo (ExpectedSig t) =
[("type", "ExpectedSig"),
("got", prettyStr True t)]
toInfo (ExpectedEq t) =
[("type", "ExpectedEq"),
("got", prettyStr True t)]
toInfo (BadUniverse k l) =
[("type", "BadUniverse"),
("low", show k),
("high", show l)]
toInfo (ClashT mode ty s t) =
[("type", "ClashT"),
("mode", show mode),
("ty", prettyStr True ty),
("left", prettyStr True s),
("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) =
[("type", "ClashU"),
("mode", show mode),
("left", prettyStr True k),
("right", prettyStr True l)]
toInfo (ClashQ pi rh) =
[("type", "ClashQ"),
("left", prettyStr True pi),
("right", prettyStr True rh)]
toInfo (ClashD p q) =
[("type", "ClashD"),
("left", prettyStr True p),
("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
M = ReaderT (Definitions Three) (Either (Error Three))
@ -68,6 +13,7 @@ defGlobals = fromList
[("A", mkAbstract Zero $ TYPE 0),
("B", mkAbstract Zero $ TYPE 0),
("a", mkAbstract Any $ FT "A"),
("a'", mkAbstract Any $ FT "A"),
("b", mkAbstract Any $ FT "B"),
("f", mkAbstract Any $ Arr One (FT "A") (FT "A"))]
@ -100,6 +46,7 @@ export
tests : Test
tests = "equality & subtyping" :- [
note #""0=1𝒥" means that 𝒥 holds in an inconsistent dim context"#,
note #""s{}" for term substs; "s" for dim substs"#,
"universes" :- [
testEq "★₀ ≡ ★₀" $
@ -117,8 +64,8 @@ tests = "equality & subtyping" :- [
],
"pi" :- [
note #""AB" for (1 _ : A) → B"#,
note #""AB" for (0 _ : A) → B"#,
note #""AB" for (1·A) → B"#,
note #""AB" for (0·A) → B"#,
testEq "A ⊸ B ≡ A ⊸ B" $
let tm = Arr One (FT "A") (FT "B") in
equalT [<] (TYPE 0) tm tm,
@ -168,32 +115,31 @@ tests = "equality & subtyping" :- [
"lambda" :- [
testEq "λ x ⇒ [x] ≡ λ x ⇒ [x]" $
equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "x" $ TUsed $ BVT 0),
(["x"] :\\ BVT 0)
(["x"] :\\ BVT 0),
testEq "λ x ⇒ [x] <: λ x ⇒ [x]" $
subT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "x" $ TUsed $ BVT 0),
(["x"] :\\ BVT 0)
(["x"] :\\ BVT 0),
testEq "λ x ⇒ [x] ≡ λ y ⇒ [y]" $
equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "y" $ TUsed $ BVT 0),
(["x"] :\\ BVT 0)
(["y"] :\\ BVT 0),
testEq "λ x ⇒ [x] <: λ y ⇒ [y]" $
equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ BVT 0)
(Lam "y" $ TUsed $ BVT 0),
(["x"] :\\ BVT 0)
(["y"] :\\ BVT 0),
testNeq "λ x y ⇒ [x] ≢ λ x y ⇒ [y]" $
equalT [<] (Arr One (FT "A") $ Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 1)
(Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 0),
(["x", "y"] :\\ BVT 1)
(["x", "y"] :\\ BVT 0),
testEq "λ x ⇒ [a] ≡ λ x ⇒ [a] (TUsed vs TUnused)" $
equalT [<] (Arr Zero (FT "B") (FT "A"))
(Lam "x" $ TUsed $ FT "a")
(Lam "x" $ TUnused $ FT "a"),
skipWith "(no η yet)" $
testEq "λ x ⇒ [f [x]] ≡ [f] (η)" $
equalT [<] (Arr One (FT "A") (FT "A"))
(Lam "x" $ TUsed $ E $ F "f" :@ BVT 0)
(["x"] :\\ E (F "f" :@ BVT 0))
(FT "f")
],
@ -208,7 +154,23 @@ tests = "equality & subtyping" :- [
(Eq0 (FT "A") (TYPE 0) (TYPE 0))
],
todo "dim lambda",
"equalities" :-
let refl : Term q d n -> Term q d n -> Elim q d n
refl a x = (DLam "_" $ DUnused x) :# (Eq0 a x x)
in
[
note #""refl [A] x" is an abbreviation for "(λᴰi ⇒ x)(x ≡ x : A)""#,
testEq "refl [A] a ≡ refl [A] a" $
equalE [<] (refl (FT "A") (FT "a")) (refl (FT "A") (FT "a")),
testEq "p : (a ≡ a' : A), q : (a ≡ a' : A) ⊢ p ≡ q (free)"
{globals =
let def = mkAbstract Zero $ Eq0 (FT "A") (FT "a") (FT "a'") in
defGlobals `mergeLeft` fromList [("p", def), ("q", def)]} $
equalE [<] (F "p") (F "q"),
testEq "x : (a ≡ a' : A), y : (a ≡ a' : A) ⊢ x ≡ y (bound)" $
let ty : forall n. Term Three 0 n := Eq0 (FT "A") (FT "a") (FT "a'") in
equalE [< ty, ty] (BV 0) (BV 1) {n = 2}
],
"term closure" :- [
note "𝑖, 𝑗 for bound variables pointing outside of the current expr",
@ -230,8 +192,8 @@ tests = "equality & subtyping" :- [
(Lam "y" $ TUnused $ FT "a"),
testEq "(λy. [𝑖]){y/y, a/𝑖} ≡ λy. [a] (TUsed)" $
equalT [<] (Arr Zero (FT "B") (FT "A"))
(CloT (Lam "y" $ TUsed $ BVT 1) (F "a" ::: id))
(Lam "y" $ TUsed $ FT "a")
(CloT (["y"] :\\ BVT 1) (F "a" ::: id))
(["y"] :\\ FT "a")
],
todo "term d-closure",
@ -290,48 +252,29 @@ tests = "equality & subtyping" :- [
subE [<] (F "f" :@ FT "a") (F "f" :@ FT "a"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ ([a ∷ A] ∷ A) (β)" $
equalE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a")
(((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
(E (FT "a" :# FT "A") :# FT "A"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a ≡ a (βυ)" $
equalE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a")
(((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
(F "a"),
testEq "(λ g ⇒ [g [a]] ∷ ⋯)) [f] ≡ (λ y ⇒ [f [y]] ∷ ⋯) [a] (β↘↙)" $
let a = FT "A"; a2a = (Arr One a a) in
equalE [<]
((Lam "g" (TUsed (E (BV 0 :@ FT "a"))) :# Arr One a2a a) :@ FT "f")
((Lam "y" (TUsed (E (F "f" :@ BVT 0))) :# a2a) :@ FT "a"),
(((["g"] :\\ E (BV 0 :@ FT "a")) :# Arr One a2a a) :@ FT "f")
(((["y"] :\\ E (F "f" :@ BVT 0)) :# a2a) :@ FT "a"),
testEq "(λ x ⇒ [x] ∷ A ⊸ A) a <: a" $
subE [<]
((Lam "x" (TUsed (BVT 0)) :# (Arr One (FT "A") (FT "A")))
:@ FT "a")
(((["x"] :\\ BVT 0) :# (Arr One (FT "A") (FT "A"))) :@ FT "a")
(F "a"),
testEq "id : A ⊸ A ≔ λ x ⇒ [x] ⊢ id [a] ≡ a"
{globals = defGlobals `mergeLeft` fromList
[("id", mkDef Any (Arr One (FT "A") (FT "A"))
(Lam "x" (TUsed (BVT 0))))]} $
(["x"] :\\ BVT 0))]} $
equalE [<] (F "id" :@ FT "a") (F "a")
],
"dim application" :-
let refl : Term q d n -> Term q d n -> Elim q d n
refl a x = (DLam "_" $ DUnused x) :# (Eq0 a x x)
in
[
note #""refl [A] x" is an abbreviation for "(λᴰi ⇒ x)(x ≡ x : A)""#,
testEq "refl [A] a ≡ refl [A] a" $
equalE [<] (refl (FT "A") (FT "a")) (refl (FT "A") (FT "a")),
testEq "p : (a ≡ b : A), q : (a ≡ b : A) ⊢ p ≡ q"
{globals =
let def = mkAbstract Zero $ Eq0 (FT "A") (FT "a") (FT "b") in
fromList [("A", mkAbstract Zero $ TYPE 0),
("a", mkAbstract Any $ FT "A"),
("b", mkAbstract Any $ FT "A"),
("p", def), ("q", def)]} $
equalE [<] (F "p") (F "q")
],
todo "dim application",
todo "annotation",

138
tests/Tests/Typechecker.idr Normal file
View file

@ -0,0 +1,138 @@
module Tests.Typechecker
import Quox.Syntax
import Quox.Syntax.Qty.Three
import Quox.Typechecker as Lib
import public TypingImpls
import TAP
0 M : Type -> Type
M = ReaderT (Definitions Three) $ Either (Error Three)
reflTy : IsQty q => Term q d n
reflTy =
Pi zero "A" (TYPE 0) $ TUsed $
Pi zero "x" (BVT 0) $ TUsed $
Eq0 (BVT 1) (BVT 0) (BVT 0)
reflDef : IsQty q => Term q d n
reflDef = ["A","x"] :\\ ["i"] :\\% BVT 0
defGlobals : Definitions Three
defGlobals = fromList
[("A", mkAbstract Zero $ TYPE 0),
("B", mkAbstract Zero $ TYPE 0),
("C", mkAbstract Zero $ TYPE 1),
("D", mkAbstract Zero $ TYPE 1),
("a", mkAbstract Any $ FT "A"),
("b", mkAbstract Any $ FT "B"),
("f", mkAbstract Any $ Arr One (FT "A") (FT "A")),
("refl", mkDef Any reflTy reflDef)]
parameters (label : String) (act : Lazy (M ()))
{default defGlobals globals : Definitions Three}
testTC : Test
testTC = test label $ runReaderT globals act
testTCFail : Test
testTCFail = testThrows label (const True) $ runReaderT globals act
ctxWith : DContext d -> Context (\i => (Term Three d i, Three)) n ->
TyContext Three d n
ctxWith dctx tqctx =
let (tctx, qctx) = unzip tqctx in MkTyContext {dctx, tctx, qctx}
ctx : Context (\i => (Term Three 0 i, Three)) n -> TyContext Three 0 n
ctx = ctxWith DNil
inferAs : TyContext Three d n -> (sg : SQty Three) ->
Elim Three d n -> Term Three d n -> M ()
inferAs ctx sg e ty = do
ty' <- infer ctx sg e
catchError
(equalType (makeDimEq ctx.dctx) ctx.tctx ty ty'.type)
(\_ : Error Three => throwError $ ClashT Equal (TYPE UAny) ty ty'.type)
infer_ : TyContext Three d n -> (sg : SQty Three) -> Elim Three d n -> M ()
infer_ ctx sg e = ignore $ infer ctx sg e
check_ : TyContext Three d n -> SQty Three ->
Term Three d n -> Term Three d n -> M ()
check_ ctx sg s ty = ignore $ check ctx sg s ty
export
tests : Test
tests = "typechecker" :- [
"universes" :- [
testTC "0 · ★₀ ⇐ ★₁" $ check_ (ctx [<]) szero (TYPE 0) (TYPE 1),
testTC "0 · ★₀ ⇐ ★₂" $ check_ (ctx [<]) szero (TYPE 0) (TYPE 2),
testTC "0 · ★₀ ⇐ ★_" $ check_ (ctx [<]) szero (TYPE 0) (TYPE UAny),
testTCFail "0 · ★₁ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 1) (TYPE 0),
testTCFail "0 · ★₀ ⇍ ★₀" $ check_ (ctx [<]) szero (TYPE 0) (TYPE 0),
testTCFail "0 · ★_ ⇍ ★_" $ check_ (ctx [<]) szero (TYPE UAny) (TYPE UAny),
testTCFail "1 · ★₀ ⇍ ★₁" $ check_ (ctx [<]) sone (TYPE 0) (TYPE 1)
],
"function types" :- [
note "A, B : ★₀; C, D : ★₁",
testTC "0 · (1·A) → B ⇐ ★₀" $
check_ (ctx [<]) szero (Arr One (FT "A") (FT "B")) (TYPE 0),
testTC "0 · (1·A) → B ⇐ ★₁👈" $
check_ (ctx [<]) szero (Arr One (FT "A") (FT "B")) (TYPE 1),
testTC "0 · (1·C) → D ⇐ ★₁" $
check_ (ctx [<]) szero (Arr One (FT "C") (FT "D")) (TYPE 1),
testTCFail "0 · (1·C) → D ⇍ ★₀" $
check_ (ctx [<]) szero (Arr One (FT "C") (FT "D")) (TYPE 0)
],
"free vars" :- [
testTC "0 · A ⇒ ★₀" $
inferAs (ctx [<]) szero (F "A") (TYPE 0),
testTC "0 · A ⇐👈 ★₀" $
check_ (ctx [<]) szero (FT "A") (TYPE 0),
testTC "0 · A ⇐ ★₁👈" $
check_ (ctx [<]) szero (FT "A") (TYPE 1),
testTCFail "1👈 · A ⇏ ★₀" $
infer_ (ctx [<]) sone (F "A"),
note "refl : (0·A : ★₀) → (0·x : A) → (x ≡ x : A) ≔ (λ A x ⇒ λᴰ _ ⇒ x)",
testTC "1 · refl ⇒ {type of refl}" $
inferAs (ctx [<]) sone (F "refl") reflTy,
testTC "1 · refl ⇐ {type of refl}" $
check_ (ctx [<]) sone (FT "refl") reflTy
],
"lambda" :- [
testTC #"1 · (λ A x ⇒ refl A x) ⇐ {type of refl, see "free vars"}"# $
check_ (ctx [<]) sone
(["A", "x"] :\\ E (F "refl" :@@ [BVT 1, BVT 0]))
reflTy
],
"misc" :- [
testTC "funext"
{globals = fromList
[("A", mkAbstract Zero $ TYPE 0),
("B", mkAbstract Zero $ Arr Any (FT "A") (TYPE 0)),
("f", mkAbstract Any $
Pi Any "x" (FT "A") $ TUsed $ E $ F "B" :@ BVT 0),
("g", mkAbstract Any $
Pi Any "x" (FT "A") $ TUsed $ E $ F "B" :@ BVT 0)]} $
-- 0·A : Type, 0·B : ω·A → Type,
-- ω·f, g : (ω·x : A) → B x
-- ⊢
-- 0·funext : (ω·eq : (0·x : A) → f x ≡ g x) → f ≡ g
-- = λ eq ⇒ λᴰ i ⇒ λ x ⇒ eq x i
check_ (ctx [<]) szero
(["eq"] :\\ ["i"] :\\% ["x"] :\\ E (BV 1 :@ BVT 0 :% BV 0))
(Pi Any "eq"
(Pi Zero "x" (FT "A") $ TUsed $
Eq0 (E $ F "B" :@ BVT 0)
(E $ F "f" :@ BVT 0) (E $ F "g" :@ BVT 0)) $ TUsed $
Eq0 (Pi Any "x" (FT "A") $ TUsed $ E $ F "B" :@ BVT 0)
(FT "f") (FT "g"))
]
]