module Tests.Reduce import Quox.Syntax as Lib import Quox.Equal import TypingImpls import AstExtra import TAP import Control.Eff %hide Prelude.App %hide Pretty.App parameters {0 isRedex : RedexTest tm} {auto _ : CanWhnf 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 <- mapFst toInfo $ runWhnf $ whnf0 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 => (BindName, 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 "1.A → B" empty $ ^Arr One (^FT "A" 0) (^FT "B" 0), testNoStep "(x: A) ⊸ B x" empty $ ^PiY One "x" (^FT "A" 0) (E $ ^App (^F "B" 0) (^BVT 0)), testNoStep "λ x ⇒ x" empty $ ^LamY "x" (^BVT 0), testNoStep "f a" empty $ E $ ^App (^F "f" 0) (^FT "a" 0) ], "neutrals" :- [ testNoStep "x" (ctx [< ("A", ^Nat)]) $ ^BV 0, testNoStep "a" empty $ ^F "a" 0, testNoStep "f a" empty $ ^App (^F "f" 0) (^FT "a" 0), testNoStep "★₀ ∷ ★₁" empty $ ^Ann (^TYPE 0) (^TYPE 1) ], "redexes" :- [ testWhnf "a ∷ A" empty (^Ann (^FT "a" 0) (^FT "A" 0)) (^F "a" 0), testWhnf "★₁ ∷ ★₃" empty (E $ ^Ann (^TYPE 1) (^TYPE 3)) (^TYPE 1), testWhnf "(λ x ⇒ x ∷ 1.A → A) a" empty (^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A" 0) (^FT "A" 0))) (^FT "a" 0)) (^F "a" 0) ], "definitions" :- [ testWhnf "a (transparent)" empty {defs = fromList [("a", ^mkDef gzero (^TYPE 1) (^TYPE 0))]} (^F "a" 0) (^Ann (^TYPE 0) (^TYPE 1)), testNoStep "a (opaque)" empty {defs = fromList [("a", ^mkPostulate gzero (^TYPE 1))]} (^F "a" 0) ], "elim closure" :- [ testWhnf "x{}" (ctx [< ("x", ^Nat)]) (CloE (Sub (^BV 0) id)) (^BV 0), testWhnf "x{a/x}" empty (CloE (Sub (^BV 0) (^F "a" 0 ::: id))) (^F "a" 0), testWhnf "x{a/y}" (ctx [< ("x", ^Nat)]) (CloE (Sub (^BV 0) (^BV 0 ::: ^F "a" 0 ::: id))) (^BV 0), testWhnf "x{(y{a/y})/x}" empty (CloE (Sub (^BV 0) ((CloE (Sub (^BV 0) (^F "a" 0 ::: id))) ::: id))) (^F "a" 0), testWhnf "(x y){f/x,a/y}" empty (CloE (Sub (^App (^BV 0) (^BVT 1)) (^F "f" 0 ::: ^F "a" 0 ::: id))) (^App (^F "f" 0) (^FT "a" 0)), testWhnf "(y ∷ x){A/x}" (ctx [< ("x", ^Nat)]) (CloE (Sub (^Ann (^BVT 1) (^BVT 0)) (^F "A" 0 ::: id))) (^BV 0), testWhnf "(y ∷ x){A/x,a/y}" empty (CloE (Sub (^Ann (^BVT 1) (^BVT 0)) (^F "A" 0 ::: ^F "a" 0 ::: id))) (^F "a" 0) ], "term closure" :- [ testWhnf "(λ y ⇒ x){a/x}" empty (CloT (Sub (^LamN (^BVT 0)) (^F "a" 0 ::: id))) (^LamN (^FT "a" 0)), testWhnf "(λy. y){a/x}" empty (CloT (Sub (^LamY "y" (^BVT 0)) (^F "a" 0 ::: id))) (^LamY "y" (^BVT 0)) ], "looking inside `E`" :- [ testWhnf "(λx. x ∷ A ⊸ A) a" empty (E $ ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A" 0) (^FT "A" 0))) (^FT "a" 0)) (^FT "a" 0) ], "nested redex" :- [ testNoStep "λ y ⇒ ((λ x ⇒ x) ∷ 1.A → A) y" empty $ ^LamY "y" (E $ ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A" 0) (^FT "A" 0))) (^BVT 0)), testNoStep "f (((λ x ⇒ x) ∷ 1.A → A) a)" empty $ ^App (^F "f" 0) (E $ ^App (^Ann (^LamY "x" (^BVT 0)) (^Arr One (^FT "A" 0) (^FT "A" 0))) (^FT "a" 0)), testNoStep "λx. (y x){x/x,a/y}" (ctx [< ("y", ^Nat)]) $ ^LamY "x" (CloT $ Sub (E $ ^App (^BV 1) (^BVT 0)) (^BV 0 ::: ^F "a" 0 ::: id)), testNoStep "f (y x){x/x,a/y}" (ctx [< ("y", ^Nat)]) $ ^App (^F "f" 0) (CloT (Sub (E $ ^App (^BV 1) (^BVT 0)) (^BV 0 ::: ^F "a" 0 ::: id))) ] ]