module Tests.Reduce import Quox.Syntax as Lib import Quox.Equal import TermImpls import TypingImpls import TAP 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} 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)] 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 "★₀" empty $ TYPE 0, testNoStep "[A] ⊸ [B]" empty $ Arr One (FT "A") (FT "B"), testNoStep "(x: [A]) ⊸ [B [x]]" empty $ Pi One (FT "A") (S [< "x"] $ Y $ E $ F "B" :@ BVT 0), testNoStep "λx. [x]" empty $ Lam $ S [< "x"] $ Y $ BVT 0, testNoStep "[f [a]]" empty $ E $ F "f" :@ FT "a" ], "neutrals" :- [ 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]" empty (FT "a" :# FT "A") (F "a"), testWhnf "[★₁ ∷ ★₃]" empty (E (TYPE 1 :# TYPE 3)) (TYPE 1), testWhnf "(λx. [x] ∷ [A] ⊸ [A]) [a]" empty ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ FT "a") (F "a") ], "definitions" :- [ testWhnf "a (transparent)" empty {defs = fromList [("a", mkDef gzero (TYPE 1) (TYPE 0))]} (F "a") (TYPE 0 :# TYPE 1) ], "elim closure" :- [ testWhnf "x{}" (ctx [< ("A", Nat)]) (CloE (BV 0) id) (BV 0), testWhnf "x{a/x}" empty (CloE (BV 0) (F "a" ::: id)) (F "a"), 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}" empty (CloE (BV 0) ((CloE (BV 0) (F "a" ::: id)) ::: id)) (F "a"), 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}" (ctx [< ("A", Nat)]) (CloE (BVT 1 :# BVT 0) (F "A" ::: id)) (BV 0), 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}" empty (CloT (Lam $ S [< "y"] $ N $ BVT 0) (F "a" ::: id)) (Lam $ S [< "y"] $ N $ FT "a"), testWhnf "(λy. y){a/x}" empty (CloT ([< "y"] :\\ BVT 0) (F "a" ::: id)) ([< "y"] :\\ BVT 0) ], "looking inside […]" :- [ 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]]" empty $ [< "y"] :\\ E ((([< "x"] :\\ BVT 0) :# Arr One (FT "A") (FT "A")) :@ BVT 0), 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}" (ctx [< ("A", Nat)]) $ [< "x"] :\\ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id), testNoStep "f ([y [x]]{x/x,a/y})" (ctx [< ("A", Nat)]) $ F "f" :@ CloT (E $ BV 1 :@ BVT 0) (BV 0 ::: F "a" ::: id) ] ]