new representation for scopes
This commit is contained in:
parent
c75f1514ba
commit
0e481a8098
14 changed files with 376 additions and 364 deletions
|
@ -141,10 +141,10 @@ tests = "equality & subtyping" :- [
|
|||
equalT empty (Arr One (FT "A") $ Arr One (FT "A") (FT "A"))
|
||||
(["x", "y"] :\\ BVT 1)
|
||||
(["x", "y"] :\\ BVT 0),
|
||||
testEq "λ x ⇒ [a] = λ x ⇒ [a] (TUsed vs TUnused)" $
|
||||
testEq "λ x ⇒ [a] = λ x ⇒ [a] (Y vs N)" $
|
||||
equalT empty (Arr Zero (FT "B") (FT "A"))
|
||||
(Lam "x" $ TUsed $ FT "a")
|
||||
(Lam "x" $ TUnused $ FT "a"),
|
||||
(Lam $ S ["x"] $ Y $ FT "a")
|
||||
(Lam $ S ["x"] $ N $ FT "a"),
|
||||
testEq "λ x ⇒ [f [x]] = [f] (η)" $
|
||||
equalT empty (Arr One (FT "A") (FT "A"))
|
||||
(["x"] :\\ E (F "f" :@ BVT 0))
|
||||
|
@ -164,7 +164,7 @@ tests = "equality & subtyping" :- [
|
|||
|
||||
"equalities and uip" :-
|
||||
let refl : Term q d n -> Term q d n -> Elim q d n
|
||||
refl a x = (DLam "_" $ DUnused x) :# (Eq0 a x x)
|
||||
refl a x = (DLam $ S ["_"] $ N x) :# (Eq0 a x x)
|
||||
in
|
||||
[
|
||||
note #""refl [A] x" is an abbreviation for "(λᴰi ⇒ x) ∷ (x ≡ x : A)""#,
|
||||
|
@ -208,7 +208,7 @@ tests = "equality & subtyping" :- [
|
|||
{globals = defGlobals `mergeLeft` fromList
|
||||
[("E", mkDef zero (TYPE 0) (Eq0 (FT "A") (FT "a") (FT "a'")))]} $
|
||||
let ty : forall n. Term Three 0 n
|
||||
:= Sig "_" (FT "E") $ TUnused $ FT "E" in
|
||||
:= Sig (FT "E") $ S ["_"] $ N $ FT "E" in
|
||||
equalE (MkTyContext new [< ty, ty]) (BV 0) (BV 1),
|
||||
|
||||
testEq "E ≔ a ≡ a' : A, F ≔ E × E ∥ x : F, y : F ⊢ x = y"
|
||||
|
@ -235,11 +235,11 @@ tests = "equality & subtyping" :- [
|
|||
equalT empty (FT "A")
|
||||
(CloT (BVT 1) (F "a" ::: F "b" ::: id))
|
||||
(FT "b"),
|
||||
testEq "(λy. [#1]){a} = λy. [a] : B ⇾ A (TUnused)" $
|
||||
testEq "(λy. [#1]){a} = λy. [a] : B ⇾ A (N)" $
|
||||
equalT empty (Arr Zero (FT "B") (FT "A"))
|
||||
(CloT (Lam "y" $ TUnused $ BVT 0) (F "a" ::: id))
|
||||
(Lam "y" $ TUnused $ FT "a"),
|
||||
testEq "(λy. [#1]){a} = λy. [a] : B ⇾ A (TUsed)" $
|
||||
(CloT (Lam $ S ["y"] $ N $ BVT 0) (F "a" ::: id))
|
||||
(Lam $ S ["y"] $ N $ FT "a"),
|
||||
testEq "(λy. [#1]){a} = λy. [a] : B ⇾ A (Y)" $
|
||||
equalT empty (Arr Zero (FT "B") (FT "A"))
|
||||
(CloT (["y"] :\\ BVT 1) (F "a" ::: id))
|
||||
(["y"] :\\ FT "a")
|
||||
|
|
|
@ -4,124 +4,113 @@ import Quox.Syntax as Lib
|
|||
import Quox.Syntax.Qty.Three
|
||||
import Quox.Equal
|
||||
import TermImpls
|
||||
import TypingImpls
|
||||
import TAP
|
||||
|
||||
|
||||
testWhnf : Eq b => Show b => (a -> (Subset b _)) -> String -> a -> b -> Test
|
||||
testWhnf whnf label from to = test "\{label} (whnf)" $ do
|
||||
let result = fst (whnf from)
|
||||
unless (result == to) $
|
||||
Left [("exp", to), ("got", result)] {a = List (String, b)}
|
||||
parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex err}
|
||||
{auto _ : ToInfo err}
|
||||
{auto _ : forall d, n. Eq (tm Three d n)}
|
||||
{auto _ : forall d, n. Show (tm Three d n)}
|
||||
{default empty defs : Definitions Three}
|
||||
{default 0 d, n : Nat}
|
||||
testWhnf : String -> tm Three d n -> tm Three d n -> Test
|
||||
testWhnf label from to = test "\{label} (whnf)" $ do
|
||||
result <- bimap toInfo fst $ whnf defs from
|
||||
unless (result == to) $ Left [("exp", show to), ("got", show result)]
|
||||
|
||||
testNoStep : Eq a => Show a => (a -> (Subset a _)) -> String -> a -> Test
|
||||
testNoStep whnf label e = test "\{label} (no step)" $ do
|
||||
let result = fst (whnf e)
|
||||
unless (result == e) $ Left [("reduced", result)] {a = List (String, a)}
|
||||
|
||||
|
||||
|
||||
parameters {default empty defs : Definitions Three} {default 0 d, n : Nat}
|
||||
testWhnfT : String -> Term Three d n -> Term Three d n -> Test
|
||||
testWhnfT = testWhnf (whnf defs)
|
||||
|
||||
testWhnfE : String -> Elim Three d n -> Elim Three d n -> Test
|
||||
testWhnfE = testWhnf (whnf defs)
|
||||
|
||||
testNoStepE : String -> Elim Three d n -> Test
|
||||
testNoStepE = testNoStep (whnf defs)
|
||||
|
||||
testNoStepT : String -> Term Three d n -> Test
|
||||
testNoStepT = testNoStep (whnf defs)
|
||||
testNoStep : String -> tm Three d n -> Test
|
||||
testNoStep label e = testWhnf label e e
|
||||
|
||||
|
||||
tests = "whnf" :- [
|
||||
"head constructors" :- [
|
||||
testNoStepT "★₀" $ TYPE 0,
|
||||
testNoStepT "[A] ⊸ [B]" $
|
||||
testNoStep "★₀" $ TYPE 0,
|
||||
testNoStep "[A] ⊸ [B]" $
|
||||
Arr One (FT "A") (FT "B"),
|
||||
testNoStepT "(x: [A]) ⊸ [B [x]]" $
|
||||
Pi One "x" (FT "A") (TUsed $ E $ F "B" :@ BVT 0),
|
||||
testNoStepT "λx. [x]" $
|
||||
Lam "x" $ TUsed $ BVT 0,
|
||||
testNoStepT "[f [a]]" $
|
||||
testNoStep "(x: [A]) ⊸ [B [x]]" $
|
||||
Pi One (FT "A") (S ["x"] $ Y $ E $ F "B" :@ BVT 0),
|
||||
testNoStep "λx. [x]" $
|
||||
Lam $ S ["x"] $ Y $ BVT 0,
|
||||
testNoStep "[f [a]]" $
|
||||
E $ F "f" :@ FT "a"
|
||||
],
|
||||
|
||||
|
||||
"neutrals" :- [
|
||||
testNoStepE "x" {n = 1} $ BV 0,
|
||||
testNoStepE "a" $ F "a",
|
||||
testNoStepE "f [a]" $ F "f" :@ FT "a",
|
||||
testNoStepE "★₀ ∷ ★₁" $ TYPE 0 :# TYPE 1
|
||||
testNoStep "x" {n = 1} $ BV 0,
|
||||
testNoStep "a" $ F "a",
|
||||
testNoStep "f [a]" $ F "f" :@ FT "a",
|
||||
testNoStep "★₀ ∷ ★₁" $ TYPE 0 :# TYPE 1
|
||||
],
|
||||
|
||||
"redexes" :- [
|
||||
testWhnfE "[a] ∷ [A]"
|
||||
testWhnf "[a] ∷ [A]"
|
||||
(FT "a" :# FT "A")
|
||||
(F "a"),
|
||||
testWhnfT "[★₁ ∷ ★₃]"
|
||||
testWhnf "[★₁ ∷ ★₃]"
|
||||
(E (TYPE 1 :# TYPE 3))
|
||||
(TYPE 1),
|
||||
testWhnfE "(λx. [x] ∷ [A] ⊸ [A]) [a]"
|
||||
((Lam "x" (TUsed $ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
testWhnf "(λx. [x] ∷ [A] ⊸ [A]) [a]"
|
||||
(((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
(F "a")
|
||||
],
|
||||
|
||||
"definitions" :- [
|
||||
testWhnfE "a (transparent)"
|
||||
testWhnf "a (transparent)"
|
||||
{defs = fromList [("a", mkDef Zero (TYPE 1) (TYPE 0))]}
|
||||
(F "a") (TYPE 0 :# TYPE 1)
|
||||
],
|
||||
|
||||
"elim closure" :- [
|
||||
testWhnfE "x{}" {n = 1}
|
||||
testWhnf "x{}" {n = 1}
|
||||
(CloE (BV 0) id)
|
||||
(BV 0),
|
||||
testWhnfE "x{a/x}"
|
||||
testWhnf "x{a/x}"
|
||||
(CloE (BV 0) (F "a" ::: id))
|
||||
(F "a"),
|
||||
testWhnfE "x{x/x,a/y}" {n = 1}
|
||||
testWhnf "x{x/x,a/y}" {n = 1}
|
||||
(CloE (BV 0) (BV 0 ::: F "a" ::: id))
|
||||
(BV 0),
|
||||
testWhnfE "x{(y{a/y})/x}"
|
||||
testWhnf "x{(y{a/y})/x}"
|
||||
(CloE (BV 0) ((CloE (BV 0) (F "a" ::: id)) ::: id))
|
||||
(F "a"),
|
||||
testWhnfE "(x y){f/x,a/y}"
|
||||
testWhnf "(x y){f/x,a/y}"
|
||||
(CloE (BV 0 :@ BVT 1) (F "f" ::: F "a" ::: id))
|
||||
(F "f" :@ FT "a"),
|
||||
testWhnfE "([y] ∷ [x]){A/x}" {n = 1}
|
||||
testWhnf "([y] ∷ [x]){A/x}" {n = 1}
|
||||
(CloE (BVT 1 :# BVT 0) (F "A" ::: id))
|
||||
(BV 0),
|
||||
testWhnfE "([y] ∷ [x]){A/x,a/y}"
|
||||
testWhnf "([y] ∷ [x]){A/x,a/y}"
|
||||
(CloE (BVT 1 :# BVT 0) (F "A" ::: F "a" ::: id))
|
||||
(F "a")
|
||||
],
|
||||
|
||||
"term closure" :- [
|
||||
testWhnfT "(λy. x){a/x}"
|
||||
(CloT (Lam "y" $ TUnused $ BVT 0) (F "a" ::: id))
|
||||
(Lam "y" $ TUnused $ FT "a"),
|
||||
testWhnfT "(λy. y){a/x}"
|
||||
(CloT (Lam "y" $ TUsed $ BVT 0) (F "a" ::: id))
|
||||
(Lam "y" $ TUsed $ BVT 0)
|
||||
testWhnf "(λy. x){a/x}"
|
||||
(CloT (Lam $ S ["y"] $ N $ BVT 0) (F "a" ::: id))
|
||||
(Lam $ S ["y"] $ N $ FT "a"),
|
||||
testWhnf "(λy. y){a/x}"
|
||||
(CloT (["y"] :\\ BVT 0) (F "a" ::: id))
|
||||
(["y"] :\\ BVT 0)
|
||||
],
|
||||
|
||||
"looking inside […]" :- [
|
||||
testWhnfT "[(λx. x ∷ A ⊸ A) [a]]"
|
||||
(E $ (Lam "x" (TUsed $ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
testWhnf "[(λx. x ∷ A ⊸ A) [a]]"
|
||||
(E $ ((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
(FT "a")
|
||||
],
|
||||
|
||||
"nested redex" :- [
|
||||
note "whnf only looks at top level redexes",
|
||||
testNoStepT "λy. [(λx. [x] ∷ [A] ⊸ [A]) [y]]" $
|
||||
Lam "y" $ TUsed $ E $
|
||||
(Lam "x" (TUsed $ BVT 0) :# Arr One (FT "A") (FT "A")) :@ BVT 0,
|
||||
testNoStepE "f [(λx. [x] ∷ [A] ⊸ [A]) [a]]" $
|
||||
testNoStep "λy. [(λx. [x] ∷ [A] ⊸ [A]) [y]]" $
|
||||
["y"] :\\ E (((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ BVT 0),
|
||||
testNoStep "f [(λx. [x] ∷ [A] ⊸ [A]) [a]]" $
|
||||
F "a" :@
|
||||
E ((Lam "x" (TUsed $ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"),
|
||||
testNoStepT "λx. [y [x]]{x/x,a/y}" {n = 1} $
|
||||
Lam "x" $ TUsed $ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id),
|
||||
testNoStepE "f ([y [x]]{x/x,a/y})" {n = 1} $
|
||||
E (((["x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"),
|
||||
testNoStep "λx. [y [x]]{x/x,a/y}" {n = 1} $
|
||||
["x"] :\\ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id),
|
||||
testNoStep "f ([y [x]]{x/x,a/y})" {n = 1} $
|
||||
F "f" :@ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id)
|
||||
]
|
||||
]
|
||||
|
|
|
@ -36,8 +36,8 @@ inj act = do
|
|||
|
||||
reflTy : IsQty q => Term q d n
|
||||
reflTy =
|
||||
Pi zero "A" (TYPE 0) $ TUsed $
|
||||
Pi one "x" (BVT 0) $ TUsed $
|
||||
Pi zero (TYPE 0) $ S ["A"] $ Y $
|
||||
Pi one (BVT 0) $ S ["x"] $ Y $
|
||||
Eq0 (BVT 1) (BVT 0) (BVT 0)
|
||||
|
||||
reflDef : IsQty q => Term q d n
|
||||
|
@ -56,8 +56,8 @@ defGlobals = fromList
|
|||
("f", mkAbstract Any $ Arr One (FT "A") (FT "A")),
|
||||
("g", mkAbstract Any $ Arr One (FT "A") (FT "B")),
|
||||
("f2", mkAbstract Any $ Arr One (FT "A") $ Arr One (FT "A") (FT "A")),
|
||||
("p", mkAbstract Any $ Pi One "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0),
|
||||
("q", mkAbstract Any $ Pi One "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0),
|
||||
("p", mkAbstract Any $ Pi One (FT "A") $ S ["x"] $ Y $ E $ F "P" :@ BVT 0),
|
||||
("q", mkAbstract Any $ Pi One (FT "A") $ S ["x"] $ Y $ E $ F "P" :@ BVT 0),
|
||||
("refl", mkDef Any reflTy reflDef)]
|
||||
|
||||
parameters (label : String) (act : Lazy (M ()))
|
||||
|
@ -139,7 +139,7 @@ tests = "typechecker" :- [
|
|||
check_ (ctx [<]) szero (Arr One (FT "C") (FT "D")) (TYPE 0),
|
||||
testTC "0 · (1·x : A) → P x ⇐ ★₀" $
|
||||
check_ (ctx [<]) szero
|
||||
(Pi One "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0)
|
||||
(Pi One (FT "A") $ S ["x"] $ Y $ E $ F "P" :@ BVT 0)
|
||||
(TYPE 0),
|
||||
testTCFail "0 · A ⊸ P ⇍ ★₀" $
|
||||
check_ (ctx [<]) szero (Arr One (FT "A") $ FT "P") (TYPE 0),
|
||||
|
@ -207,21 +207,20 @@ tests = "typechecker" :- [
|
|||
|
||||
"equalities" :- [
|
||||
testTC "1 · (λᴰ i ⇒ a) ⇐ a ≡ a" $
|
||||
check_ (ctx [<]) sone (DLam "i" $ DUnused $ FT "a")
|
||||
check_ (ctx [<]) sone (DLam $ S ["i"] $ N $ FT "a")
|
||||
(Eq0 (FT "A") (FT "a") (FT "a")),
|
||||
testTC "0 · (λ p q ⇒ λᴰ i ⇒ p) ⇐ (ω·p q : a ≡ a') → p ≡ q" $
|
||||
check_ (ctx [<]) szero
|
||||
(Lam "p" $ TUsed $ Lam "q" $ TUnused $
|
||||
DLam "i" $ DUnused $ BVT 0)
|
||||
(Pi Any "p" (Eq0 (FT "A") (FT "a") (FT "a")) $ TUsed $
|
||||
Pi Any "q" (Eq0 (FT "A") (FT "a") (FT "a")) $ TUsed $
|
||||
(Lam $ S ["p"] $ Y $ Lam $ S ["q"] $ N $ DLam $ S ["i"] $ N $ BVT 0)
|
||||
(Pi Any (Eq0 (FT "A") (FT "a") (FT "a")) $ S ["p"] $ Y $
|
||||
Pi Any (Eq0 (FT "A") (FT "a") (FT "a")) $ S ["q"] $ Y $
|
||||
Eq0 (Eq0 (FT "A") (FT "a") (FT "a")) (BVT 1) (BVT 0)),
|
||||
testTC "0 · (λ p q ⇒ λᴰ i ⇒ q) ⇐ (ω·p q : a ≡ a') → p ≡ q" $
|
||||
check_ (ctx [<]) szero
|
||||
(Lam "p" $ TUnused $ Lam "q" $ TUsed $
|
||||
DLam "i" $ DUnused $ BVT 0)
|
||||
(Pi Any "p" (Eq0 (FT "A") (FT "a") (FT "a")) $ TUsed $
|
||||
Pi Any "q" (Eq0 (FT "A") (FT "a") (FT "a")) $ TUsed $
|
||||
(Lam $ S ["p"] $ N $ Lam $ S ["q"] $ Y $
|
||||
DLam $ S ["i"] $ N $ BVT 0)
|
||||
(Pi Any (Eq0 (FT "A") (FT "a") (FT "a")) $ S ["p"] $ Y $
|
||||
Pi Any (Eq0 (FT "A") (FT "a") (FT "a")) $ S ["q"] $ Y $
|
||||
Eq0 (Eq0 (FT "A") (FT "a") (FT "a")) (BVT 1) (BVT 0))
|
||||
],
|
||||
|
||||
|
@ -233,11 +232,11 @@ tests = "typechecker" :- [
|
|||
testTC "cong" $
|
||||
check_ (ctx [<]) sone
|
||||
(["x", "y", "xy"] :\\ ["i"] :\\% E (F "p" :@ E (BV 0 :% BV 0)))
|
||||
(Pi Zero "x" (FT "A") $ TUsed $
|
||||
Pi Zero "y" (FT "A") $ TUsed $
|
||||
Pi One "xy" (Eq0 (FT "A") (BVT 1) (BVT 0)) $ TUsed $
|
||||
Eq "i" (DUsed $ E $ F "P" :@ E (BV 0 :% BV 0))
|
||||
(E $ F "p" :@ BVT 2) (E $ F "p" :@ BVT 1)),
|
||||
(Pi Zero (FT "A") $ S ["x"] $ Y $
|
||||
Pi Zero (FT "A") $ S ["y"] $ Y $
|
||||
Pi One (Eq0 (FT "A") (BVT 1) (BVT 0)) $ S ["xy"] $ Y $
|
||||
Eq (S ["i"] $ Y $ E $ F "P" :@ E (BV 0 :% BV 0))
|
||||
(E $ F "p" :@ BVT 2) (E $ F "p" :@ BVT 1)),
|
||||
note "0·A : Type, 0·P : ω·A → Type,",
|
||||
note "ω·p q : (1·x : A) → P x",
|
||||
note "⊢",
|
||||
|
@ -246,11 +245,12 @@ tests = "typechecker" :- [
|
|||
testTC "funext" $
|
||||
check_ (ctx [<]) sone
|
||||
(["eq"] :\\ ["i"] :\\% ["x"] :\\ E (BV 1 :@ BVT 0 :% BV 0))
|
||||
(Pi One "eq"
|
||||
(Pi One "x" (FT "A") $ TUsed $
|
||||
(Pi One
|
||||
(Pi One (FT "A") $ S ["x"] $ Y $
|
||||
Eq0 (E $ F "P" :@ BVT 0)
|
||||
(E $ F "p" :@ BVT 0) (E $ F "q" :@ BVT 0)) $ TUsed $
|
||||
Eq0 (Pi Any "x" (FT "A") $ TUsed $ E $ F "P" :@ BVT 0)
|
||||
(FT "p") (FT "q"))
|
||||
(E $ F "p" :@ BVT 0) (E $ F "q" :@ BVT 0)) $
|
||||
S ["eq"] $ Y $
|
||||
Eq0 (Pi Any (FT "A") $ S ["x"] $ Y $ E $ F "P" :@ BVT 0)
|
||||
(FT "p") (FT "q"))
|
||||
]
|
||||
]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue