some reduction tests & fixes

This commit is contained in:
rhiannon morris 2022-05-25 16:10:19 +02:00
parent bd57a976a8
commit bc9344c6ba
4 changed files with 154 additions and 17 deletions

View File

@ -83,7 +83,10 @@ mutual
pushSubstsEWith th ph (F x) =
ncloE $ F x
pushSubstsEWith th ph (B i) =
assert_total pushSubstsE $ ph !! i
let res = ph !! i in
case choose $ topCloE res of
Left _ => assert_total pushSubstsE res
Right _ => ncloE res
pushSubstsEWith th ph (f :@ s) =
ncloE $ subs f th ph :@ subs s th ph
pushSubstsEWith th ph (s :# a) =

View File

@ -5,6 +5,7 @@ import TAP
import Tests.Unicode
import Tests.Lexer
import Tests.Parser
import Tests.Reduce
import Tests.Equal
import System
@ -13,6 +14,7 @@ allTests = [
Unicode.tests,
Lexer.tests,
Parser.tests,
Reduce.tests,
Equal.tests
]

View File

@ -1,6 +1,6 @@
module Tests.Equal
import Quox.Equal
import Quox.Equal as Lib
import Quox.Pretty
import TAP
@ -32,20 +32,20 @@ testNeq label = testThrows label $ const True
subT : {default 0 d, n : Nat} -> Term d n -> Term d n -> M ()
subT = Quox.Equal.subT
%hide Quox.Equal.subT
subT = Lib.subT
%hide Lib.subT
equalT : {default 0 d, n : Nat} -> Term d n -> Term d n -> M ()
equalT = Quox.Equal.equalT
%hide Quox.Equal.equalT
equalT = Lib.equalT
%hide Lib.equalT
subE : {default 0 d, n : Nat} -> Elim d n -> Elim d n -> M ()
subE = Quox.Equal.subE
%hide Quox.Equal.subE
subE = Lib.subE
%hide Lib.subE
equalE : {default 0 d, n : Nat} -> Elim d n -> Elim d n -> M ()
equalE = Quox.Equal.equalE
%hide Quox.Equal.equalE
equalE = Lib.equalE
%hide Lib.equalE
export
@ -112,19 +112,35 @@ tests = "equality & subtyping" :- [
"lambda" :- [
testEq "λ x ⇒ [x] ≡ λ x ⇒ [x]" $
equalT (Lam "x" (TUsed (BVT 0))) (Lam "x" (TUsed (BVT 0))),
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "x" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] <: λ x ⇒ [x]" $
equalT (Lam "x" (TUsed (BVT 0))) (Lam "x" (TUsed (BVT 0))),
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "x" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] ≡ λ y ⇒ [y]" $
equalT (Lam "x" (TUsed (BVT 0))) (Lam "y" (TUsed (BVT 0))),
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "y" $ TUsed $ BVT 0),
testEq "λ x ⇒ [x] <: λ y ⇒ [y]" $
equalT (Lam "x" (TUsed (BVT 0))) (Lam "y" (TUsed (BVT 0))),
equalT (Lam "x" $ TUsed $ BVT 0) (Lam "y" $ TUsed $ BVT 0),
testNeq "λ x y ⇒ [x] ≢ λ x y ⇒ [y]" $
equalT (Lam "x" (TUsed (Lam "y" (TUsed (BVT 1)))))
(Lam "x" (TUsed (Lam "y" (TUsed (BVT 0)))))
equalT (Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 1)
(Lam "x" $ TUsed $ Lam "y" $ TUsed $ BVT 0),
testEq "λ x ⇒ [a] ≡ λ x ⇒ [a] (TUsed vs TUnused)" $
equalT (Lam "x" $ TUsed $ FT "a")
(Lam "x" $ TUnused $ FT "a")
],
todo "term closure",
"term closure" :- [
testEq "[x]{} ≡ [x]" $
equalT (CloT (BVT 0) id) (BVT 0) {n = 1},
testEq "[x]{a/x} ≡ [a]" $
equalT (CloT (BVT 0) (F "a" ::: id)) (FT "a"),
testEq "[x]{a/x,b/y} ≡ [a]" $
equalT (CloT (BVT 0) (F "a" ::: F "b" ::: id)) (FT "a"),
testEq "(λy. [x]){y/y, a/x} ≡ λy. [a] (TUnused)" $
equalT (CloT (Lam "y" $ TUnused $ BVT 0) (F "a" ::: id))
(Lam "y" $ TUnused $ FT "a"),
testEq "(λy. [x]){y/y, a/x} ≡ λy. [a] (TUsed)" $
equalT (CloT (Lam "y" $ TUsed $ BVT 1) (F "a" ::: id))
(Lam "y" $ TUsed $ FT "a")
],
todo "term d-closure",

116
tests/Tests/Reduce.idr Normal file
View File

@ -0,0 +1,116 @@
module Tests.Reduce
import Quox.Syntax as Lib
import TermImpls
import TAP
testWhnf : (Eq b, Show b) => (a -> (Subset b _)) ->
String -> a -> b -> Test
testWhnf whnf label from to = test "\{label} (whnf)" $
let result = fst (whnf from) in
if result == to
then Right ()
else with Prelude.(::)
Left [("expected", to), ("received", result)]
testNoStep : forall p. Show a => ((x : a) -> Either (p x) a) ->
String -> a -> Test
testNoStep step label e = test "\{label} (no step)" $
case step e of
Left _ => Right ()
Right e' => with Prelude.(::) Left [("reduced", e')]
parameters {default 0 d, n : Nat}
testWhnfT : String -> Term d n -> Term d n -> Test
testWhnfT = testWhnf whnfT
testWhnfE : String -> Elim d n -> Elim d n -> Test
testWhnfE = testWhnf whnfE
testNoStepE : String -> Elim d n -> Test
testNoStepE = testNoStep stepE
testNoStepT : String -> Term d n -> Test
testNoStepT = testNoStep stepT
tests = "whnf" :- [
"head constructors" :- [
testNoStepT "★₀" $ TYPE 0,
testNoStepT "[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]]" $
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
],
"redexes" :- [
testWhnfE "[a] ∷ [A]"
(FT "a" :# FT "A")
(F "a"),
testWhnfT "[★₁ ∷ ★₃]"
(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")
(F "a")
],
"elim closure" :- [
testWhnfE "x{}" {n = 1}
(CloE (BV 0) id)
(BV 0),
testWhnfE "x{a/x}"
(CloE (BV 0) (F "a" ::: id))
(F "a"),
testWhnfE "x{x/x,a/y}" {n = 1}
(CloE (BV 0) (BV 0 ::: F "a" ::: id))
(BV 0),
testWhnfE "x{(y{a/y})/x}"
(CloE (BV 0) ((CloE (BV 0) (F "a" ::: id)) ::: id))
(F "a"),
testWhnfE "(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}
(CloE (BVT 1 :# BVT 0) (F "A" ::: id))
(BV 0),
testWhnfE "([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 "x" $ TUnused $ BVT 0) (F "a" ::: id))
(Lam "x" $ TUnused $ FT "a")
],
"looking inside […]" :- [
testWhnfT "[(λx. x ∷ A ⊸ A) [a]]"
(E $ (Lam "x" (TUsed $ 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]]" $
F "a" :@
E ((Lam "x" (TUsed $ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a")
]
]