coercions and compositions
This commit is contained in:
parent
468ae7e444
commit
a5ccf0215a
25 changed files with 1344 additions and 651 deletions
|
@ -25,14 +25,14 @@ mutual
|
|||
TYPE _ == _ = False
|
||||
|
||||
Pi qty1 arg1 res1 == Pi qty2 arg2 res2 =
|
||||
qty1 == qty2 && arg1 == arg2 && res1.term == res2.term
|
||||
qty1 == qty2 && arg1 == arg2 && res1 == res2
|
||||
Pi {} == _ = False
|
||||
|
||||
Lam body1 == Lam body2 = body1.term == body2.term
|
||||
Lam body1 == Lam body2 = body1 == body2
|
||||
Lam {} == _ = False
|
||||
|
||||
Sig fst1 snd1 == Sig fst2 snd2 =
|
||||
fst1 == fst2 && snd1.term == snd2.term
|
||||
fst1 == fst2 && snd1 == snd2
|
||||
Sig {} == _ = False
|
||||
|
||||
Pair fst1 snd1 == Pair fst2 snd2 = fst1 == fst2 && snd1 == snd2
|
||||
|
@ -45,10 +45,10 @@ mutual
|
|||
Tag _ == _ = False
|
||||
|
||||
Eq ty1 l1 r1 == Eq ty2 l2 r2 =
|
||||
ty1.term == ty2.term && l1 == l2 && r1 == r2
|
||||
ty1 == ty2 && l1 == l2 && r1 == r2
|
||||
Eq {} == _ = False
|
||||
|
||||
DLam body1 == DLam body2 = body1.term == body2.term
|
||||
DLam body1 == DLam body2 = body1 == body2
|
||||
DLam {} == _ = False
|
||||
|
||||
Nat == Nat = True
|
||||
|
@ -93,20 +93,20 @@ mutual
|
|||
(_ :@ _) == _ = False
|
||||
|
||||
CasePair q1 p1 r1 b1 == CasePair q2 p2 r2 b2 =
|
||||
q1 == q2 && p1 == p2 && r1.term == r2.term && b1.term == b2.term
|
||||
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.term == r2.term && a1 == 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.term == r2.term && z1 == z2 && s1.term == s2.term
|
||||
r1 == r2 && z1 == z2 && s1 == s2
|
||||
CaseNat {} == _ = False
|
||||
|
||||
CaseBox q1 x1 r1 b1 == CaseBox q2 x2 r2 b2 =
|
||||
q1 == q2 && x1 == x2 && r1.term == r2.term && b1.term == b2.term
|
||||
q1 == q2 && x1 == x2 && r1 == r2 && b1 == b2
|
||||
CaseBox {} == _ = False
|
||||
|
||||
(fun1 :% dim1) == (fun2 :% dim2) = fun1 == fun2 && dim1 == dim2
|
||||
|
@ -115,17 +115,18 @@ mutual
|
|||
(tm1 :# ty1) == (tm2 :# ty2) = tm1 == tm2 && ty1 == ty2
|
||||
(_ :# _) == _ = False
|
||||
|
||||
TypeCase ty1 ret1 univ1 pi1 sig1 enum1 eq1 nat1 box1
|
||||
==
|
||||
TypeCase ty2 ret2 univ2 pi2 sig2 enum2 eq2 nat2 box2 =
|
||||
ty1 == ty2 && ret1 == ret2 &&
|
||||
pi1.term == pi2.term &&
|
||||
sig1.term == sig2.term &&
|
||||
enum1 == enum2 &&
|
||||
eq1.term == eq2.term &&
|
||||
nat1 == nat2 &&
|
||||
box1.term == box2.term
|
||||
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 =
|
||||
|
@ -140,6 +141,14 @@ mutual
|
|||
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
|
||||
|
|
|
@ -155,30 +155,9 @@ tests = "pretty printing terms" :- [
|
|||
"type-case" :- [
|
||||
testPrettyE [<] [<]
|
||||
{label = "type-case ℕ ∷ ★₀ return ★₀ of { ⋯ }"}
|
||||
(TypeCase (Nat :# TYPE 0) (TYPE 0) Nat (SN Nat) (SN Nat) Nat
|
||||
(SN Nat) Nat (SN Nat))
|
||||
"""
|
||||
type-case ℕ ∷ ★₀ return ★₀ of {
|
||||
★ ⇒ ℕ;
|
||||
(_ → _) ⇒ ℕ;
|
||||
(_ × _) ⇒ ℕ;
|
||||
{} ⇒ ℕ;
|
||||
Eq _ _ _ _ _ ⇒ ℕ;
|
||||
ℕ ⇒ ℕ;
|
||||
[_] ⇒ ℕ
|
||||
}
|
||||
"""
|
||||
"""
|
||||
type-case Nat :: Type0 return Type0 of {
|
||||
Type => Nat;
|
||||
(_ -> _) => Nat;
|
||||
(_ ** _) => Nat;
|
||||
{} => Nat;
|
||||
Eq _ _ _ _ _ => Nat;
|
||||
Nat => Nat;
|
||||
[_] => Nat
|
||||
}
|
||||
"""
|
||||
(TypeCase (Nat :# TYPE 0) (TYPE 0) empty Nat)
|
||||
"type-case ℕ ∷ ★₀ return ★₀ of { _ ⇒ ℕ }"
|
||||
"type-case Nat :: Type0 return Type0 of { _ => Nat }"
|
||||
],
|
||||
|
||||
"annotations" :- [
|
||||
|
|
|
@ -7,108 +7,113 @@ import TypingImpls
|
|||
import TAP
|
||||
|
||||
|
||||
parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex err}
|
||||
{auto _ : ToInfo err}
|
||||
{auto _ : forall d, n. Eq (tm d n)}
|
||||
{auto _ : forall d, n. Show (tm d n)}
|
||||
parameters {0 isRedex : RedexTest tm} {auto _ : Whnf tm isRedex} {d, n : Nat}
|
||||
{auto _ : (Eq (tm d n), Show (tm d n))}
|
||||
{default empty defs : Definitions}
|
||||
{default 0 d, n : Nat}
|
||||
testWhnf : String -> tm d n -> tm d n -> Test
|
||||
testWhnf label from to = test "\{label} (whnf)" $ do
|
||||
result <- bimap toInfo fst $ whnf defs from
|
||||
private
|
||||
testWhnf : String -> WhnfContext d n -> tm d n -> tm d n -> Test
|
||||
testWhnf label ctx from to = test "\{label} (whnf)" $ do
|
||||
result <- bimap toInfo fst $ whnf defs ctx from
|
||||
unless (result == to) $ Left [("exp", show to), ("got", show result)]
|
||||
|
||||
testNoStep : String -> tm d n -> Test
|
||||
testNoStep label e = testWhnf label e e
|
||||
private
|
||||
testNoStep : String -> WhnfContext d n -> tm d n -> Test
|
||||
testNoStep label ctx e = testWhnf label ctx e e
|
||||
|
||||
private
|
||||
ctx : Context (\n => (BaseName, Term 0 n)) n -> WhnfContext 0 n
|
||||
ctx xs = let (ns, ts) = unzip xs in MkWhnfContext [<] ns ts
|
||||
|
||||
|
||||
export
|
||||
tests : Test
|
||||
tests = "whnf" :- [
|
||||
"head constructors" :- [
|
||||
testNoStep "★₀" $ TYPE 0,
|
||||
testNoStep "[A] ⊸ [B]" $
|
||||
testNoStep "★₀" empty $ TYPE 0,
|
||||
testNoStep "[A] ⊸ [B]" empty $
|
||||
Arr One (FT "A") (FT "B"),
|
||||
testNoStep "(x: [A]) ⊸ [B [x]]" $
|
||||
testNoStep "(x: [A]) ⊸ [B [x]]" empty $
|
||||
Pi One (FT "A") (S [< "x"] $ Y $ E $ F "B" :@ BVT 0),
|
||||
testNoStep "λx. [x]" $
|
||||
testNoStep "λx. [x]" empty $
|
||||
Lam $ S [< "x"] $ Y $ BVT 0,
|
||||
testNoStep "[f [a]]" $
|
||||
testNoStep "[f [a]]" empty $
|
||||
E $ F "f" :@ FT "a"
|
||||
],
|
||||
|
||||
"neutrals" :- [
|
||||
testNoStep "x" {n = 1} $ BV 0,
|
||||
testNoStep "a" $ F "a",
|
||||
testNoStep "f [a]" $ F "f" :@ FT "a",
|
||||
testNoStep "★₀ ∷ ★₁" $ TYPE 0 :# TYPE 1
|
||||
testNoStep "x" (ctx [< ("A", Nat)]) $ BV 0,
|
||||
testNoStep "a" empty $ F "a",
|
||||
testNoStep "f [a]" empty $ F "f" :@ FT "a",
|
||||
testNoStep "★₀ ∷ ★₁" empty $ TYPE 0 :# TYPE 1
|
||||
],
|
||||
|
||||
"redexes" :- [
|
||||
testWhnf "[a] ∷ [A]"
|
||||
testWhnf "[a] ∷ [A]" empty
|
||||
(FT "a" :# FT "A")
|
||||
(F "a"),
|
||||
testWhnf "[★₁ ∷ ★₃]"
|
||||
testWhnf "[★₁ ∷ ★₃]" empty
|
||||
(E (TYPE 1 :# TYPE 3))
|
||||
(TYPE 1),
|
||||
testWhnf "(λx. [x] ∷ [A] ⊸ [A]) [a]"
|
||||
testWhnf "(λx. [x] ∷ [A] ⊸ [A]) [a]" empty
|
||||
((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
(F "a")
|
||||
],
|
||||
|
||||
"definitions" :- [
|
||||
testWhnf "a (transparent)"
|
||||
testWhnf "a (transparent)" empty
|
||||
{defs = fromList [("a", mkDef gzero (TYPE 1) (TYPE 0))]}
|
||||
(F "a") (TYPE 0 :# TYPE 1)
|
||||
],
|
||||
|
||||
"elim closure" :- [
|
||||
testWhnf "x{}" {n = 1}
|
||||
testWhnf "x{}" (ctx [< ("A", Nat)])
|
||||
(CloE (BV 0) id)
|
||||
(BV 0),
|
||||
testWhnf "x{a/x}"
|
||||
testWhnf "x{a/x}" empty
|
||||
(CloE (BV 0) (F "a" ::: id))
|
||||
(F "a"),
|
||||
testWhnf "x{x/x,a/y}" {n = 1}
|
||||
testWhnf "x{x/x,a/y}" (ctx [< ("A", Nat)])
|
||||
(CloE (BV 0) (BV 0 ::: F "a" ::: id))
|
||||
(BV 0),
|
||||
testWhnf "x{(y{a/y})/x}"
|
||||
testWhnf "x{(y{a/y})/x}" empty
|
||||
(CloE (BV 0) ((CloE (BV 0) (F "a" ::: id)) ::: id))
|
||||
(F "a"),
|
||||
testWhnf "(x y){f/x,a/y}"
|
||||
testWhnf "(x y){f/x,a/y}" empty
|
||||
(CloE (BV 0 :@ BVT 1) (F "f" ::: F "a" ::: id))
|
||||
(F "f" :@ FT "a"),
|
||||
testWhnf "([y] ∷ [x]){A/x}" {n = 1}
|
||||
testWhnf "([y] ∷ [x]){A/x}" (ctx [< ("A", Nat)])
|
||||
(CloE (BVT 1 :# BVT 0) (F "A" ::: id))
|
||||
(BV 0),
|
||||
testWhnf "([y] ∷ [x]){A/x,a/y}"
|
||||
testWhnf "([y] ∷ [x]){A/x,a/y}" empty
|
||||
(CloE (BVT 1 :# BVT 0) (F "A" ::: F "a" ::: id))
|
||||
(F "a")
|
||||
],
|
||||
|
||||
"term closure" :- [
|
||||
testWhnf "(λy. x){a/x}"
|
||||
testWhnf "(λy. x){a/x}" empty
|
||||
(CloT (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id))
|
||||
(Lam $ S [< "y"] $ N $ FT "a"),
|
||||
testWhnf "(λy. y){a/x}"
|
||||
testWhnf "(λy. y){a/x}" empty
|
||||
(CloT ([< "y"] :\\ BVT 0) (F "a" ::: id))
|
||||
([< "y"] :\\ BVT 0)
|
||||
],
|
||||
|
||||
"looking inside […]" :- [
|
||||
testWhnf "[(λx. x ∷ A ⊸ A) [a]]"
|
||||
testWhnf "[(λx. x ∷ A ⊸ A) [a]]" empty
|
||||
(E $ (([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
|
||||
(FT "a")
|
||||
],
|
||||
|
||||
"nested redex" :- [
|
||||
note "whnf only looks at top level redexes",
|
||||
testNoStep "λy. [(λx. [x] ∷ [A] ⊸ [A]) [y]]" $
|
||||
testNoStep "λy. [(λx. [x] ∷ [A] ⊸ [A]) [y]]" empty $
|
||||
[< "y"] :\\ E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ BVT 0),
|
||||
testNoStep "f [(λx. [x] ∷ [A] ⊸ [A]) [a]]" $
|
||||
testNoStep "f [(λx. [x] ∷ [A] ⊸ [A]) [a]]" empty $
|
||||
F "a" :@
|
||||
E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a"),
|
||||
testNoStep "λx. [y [x]]{x/x,a/y}" {n = 1} $
|
||||
testNoStep "λx. [y [x]]{x/x,a/y}" (ctx [< ("A", Nat)]) $
|
||||
[< "x"] :\\ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id),
|
||||
testNoStep "f ([y [x]]{x/x,a/y})" {n = 1} $
|
||||
testNoStep "f ([y [x]]{x/x,a/y})" (ctx [< ("A", Nat)]) $
|
||||
F "f" :@ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id)
|
||||
]
|
||||
]
|
||||
|
|
|
@ -451,8 +451,7 @@ tests = "typechecker" :- [
|
|||
"type-case" :- [
|
||||
testTC "0 · type-case ℕ ∷ ★₀ return ★₀ of { _ ⇒ ℕ } ⇒ ★₀" $
|
||||
inferAs empty szero
|
||||
(TypeCase (Nat :# TYPE 0) (TYPE 0) Nat (SN Nat) (SN Nat) Nat
|
||||
(SN Nat) Nat (SN Nat))
|
||||
(TypeCase (Nat :# TYPE 0) (TYPE 0) empty Nat)
|
||||
(TYPE 0)
|
||||
],
|
||||
|
||||
|
|
|
@ -9,17 +9,10 @@ import Derive.Prelude
|
|||
%language ElabReflection
|
||||
|
||||
|
||||
%runElab derive "Reduce.WhnfError" [Show]
|
||||
%runElab deriveIndexed "TyContext" [Show]
|
||||
%runElab deriveIndexed "EqContext" [Show]
|
||||
%runElab deriveIndexed "NameContexts" [Show]
|
||||
%runElab derive "Error" [Show]
|
||||
|
||||
export
|
||||
ToInfo WhnfError where
|
||||
toInfo (MissingEnumArm t ts) =
|
||||
[("type", "MissingEnumArm"),
|
||||
("tag", show t),
|
||||
("list", show ts)]
|
||||
|
||||
export
|
||||
ToInfo Error where toInfo err = [("err", show $ prettyError True True err)]
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue