fix up tests
This commit is contained in:
parent
5e220da2f4
commit
a17752f31c
5 changed files with 38 additions and 35 deletions
|
@ -13,7 +13,7 @@ import System.File
|
||||||
import System.Path
|
import System.Path
|
||||||
import Data.IORef
|
import Data.IORef
|
||||||
|
|
||||||
import public Quox.Parser.FromParser.Error
|
import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
|
||||||
|
|
||||||
%default total
|
%default total
|
||||||
|
|
||||||
|
|
|
@ -1,12 +1,14 @@
|
||||||
module Tests.Equal
|
module Tests.Equal
|
||||||
|
|
||||||
import Quox.Equal
|
import Quox.Equal
|
||||||
|
import Quox.Typechecker
|
||||||
import Quox.Syntax.Qty.Three
|
import Quox.Syntax.Qty.Three
|
||||||
import public TypingImpls
|
import public TypingImpls
|
||||||
import TAP
|
import TAP
|
||||||
|
import Quox.EffExtra
|
||||||
|
|
||||||
0 M : Type -> Type
|
0 M : Type -> Type
|
||||||
M = ReaderT (Definitions Three) (Either (Error Three))
|
M = TC Three
|
||||||
|
|
||||||
defGlobals : Definitions Three
|
defGlobals : Definitions Three
|
||||||
defGlobals = fromList
|
defGlobals = fromList
|
||||||
|
@ -23,10 +25,10 @@ defGlobals = fromList
|
||||||
parameters (label : String) (act : Lazy (M ()))
|
parameters (label : String) (act : Lazy (M ()))
|
||||||
{default defGlobals globals : Definitions Three}
|
{default defGlobals globals : Definitions Three}
|
||||||
testEq : Test
|
testEq : Test
|
||||||
testEq = test label $ runReaderT globals act
|
testEq = test label $ runTC globals act
|
||||||
|
|
||||||
testNeq : Test
|
testNeq : Test
|
||||||
testNeq = testThrows label (const True) $ runReaderT globals act $> "()"
|
testNeq = testThrows label (const True) $ runTC globals act $> "()"
|
||||||
|
|
||||||
|
|
||||||
parameters (0 d : Nat) (ctx : TyContext Three d n)
|
parameters (0 d : Nat) (ctx : TyContext Three d n)
|
||||||
|
|
|
@ -1,11 +1,12 @@
|
||||||
module Tests.FromPTerm
|
module Tests.FromPTerm
|
||||||
|
|
||||||
import Quox.Parser.Syntax
|
import Quox.Parser.FromParser
|
||||||
import Quox.Parser
|
import Quox.Parser
|
||||||
import TermImpls
|
import TermImpls
|
||||||
import TypingImpls
|
import TypingImpls
|
||||||
import Tests.Parser as TParser
|
import Tests.Parser as TParser
|
||||||
import TAP
|
import TAP
|
||||||
|
import Quox.EffExtra
|
||||||
|
|
||||||
import System.File
|
import System.File
|
||||||
import Derive.Prelude
|
import Derive.Prelude
|
||||||
|
@ -16,21 +17,19 @@ import Derive.Prelude
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data Failure =
|
data Failure =
|
||||||
ParseError (Parser.Error)
|
ParseError Parser.Error
|
||||||
| FromParser FromParserError
|
| FromParser FromParser.Error
|
||||||
| WrongResult String
|
| WrongResult String
|
||||||
| ExpectedFail String
|
| ExpectedFail String
|
||||||
|
|
||||||
%runElab derive "FileError" [Show]
|
%runElab derive "FileError" [Show]
|
||||||
%runElab derive "Parser.Error" [Show]
|
|
||||||
%runElab derive "FromParserError" [Show]
|
|
||||||
|
|
||||||
export
|
export
|
||||||
ToInfo Failure where
|
ToInfo Failure where
|
||||||
toInfo (ParseError err) = toInfo err
|
toInfo (ParseError err) = toInfo err
|
||||||
toInfo (FromParser err) =
|
toInfo (FromParser err) =
|
||||||
[("type", "FromParserError"),
|
[("type", "FromParserError"),
|
||||||
("got", show err)]
|
("got", show $ prettyError True True err)]
|
||||||
toInfo (WrongResult got) =
|
toInfo (WrongResult got) =
|
||||||
[("type", "WrongResult"), ("got", got)]
|
[("type", "WrongResult"), ("got", got)]
|
||||||
toInfo (ExpectedFail got) =
|
toInfo (ExpectedFail got) =
|
||||||
|
@ -38,7 +37,7 @@ ToInfo Failure where
|
||||||
|
|
||||||
|
|
||||||
parameters {c : Bool} {auto _ : Show b}
|
parameters {c : Bool} {auto _ : Show b}
|
||||||
(grm : Grammar c a) (fromP : a -> Either FromParserError b)
|
(grm : Grammar c a) (fromP : a -> Either FromParser.Error b)
|
||||||
(inp : String)
|
(inp : String)
|
||||||
parameters {default (ltrim inp) label : String}
|
parameters {default (ltrim inp) label : String}
|
||||||
parsesWith : (b -> Bool) -> Test
|
parsesWith : (b -> Bool) -> Test
|
||||||
|
@ -62,11 +61,14 @@ parameters {c : Bool} {auto _ : Show b}
|
||||||
|
|
||||||
FromString BName where fromString = Just . fromString
|
FromString BName where fromString = Just . fromString
|
||||||
|
|
||||||
|
runFromParser : Eff [Except FromParser.Error] a -> Either FromParser.Error a
|
||||||
|
runFromParser = extract . runExcept
|
||||||
|
|
||||||
export
|
export
|
||||||
tests : Test
|
tests : Test
|
||||||
tests = "PTerm → Term" :- [
|
tests = "PTerm → Term" :- [
|
||||||
"dimensions" :-
|
"dimensions" :-
|
||||||
let fromPDim = fromPDimWith [< "𝑖", "𝑗"]
|
let fromPDim = runFromParser . fromPDimWith [< "𝑖", "𝑗"]
|
||||||
in [
|
in [
|
||||||
note "dim ctx: [𝑖, 𝑗]",
|
note "dim ctx: [𝑖, 𝑗]",
|
||||||
parsesAs dim fromPDim "𝑖" (BV 1),
|
parsesAs dim fromPDim "𝑖" (BV 1),
|
||||||
|
@ -77,18 +79,19 @@ tests = "PTerm → Term" :- [
|
||||||
],
|
],
|
||||||
|
|
||||||
"terms" :-
|
"terms" :-
|
||||||
let fromPTerm' = fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
|
let fromPTerm = runFromParser .
|
||||||
|
fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
|
||||||
in [
|
in [
|
||||||
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",
|
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",
|
||||||
parsesAs term fromPTerm' "x" $ BVT 2,
|
parsesAs term fromPTerm "x" $ BVT 2,
|
||||||
parseFails term fromPTerm' "𝑖",
|
parseFails term fromPTerm "𝑖",
|
||||||
parsesAs term fromPTerm' "f" $ FT "f",
|
parsesAs term fromPTerm "f" $ FT "f",
|
||||||
parsesAs term fromPTerm' "λ w ⇒ w" $ [< "w"] :\\ BVT 0,
|
parsesAs term fromPTerm "λ w ⇒ w" $ [< "w"] :\\ BVT 0,
|
||||||
parsesAs term fromPTerm' "λ w ⇒ x" $ [< "w"] :\\ BVT 3,
|
parsesAs term fromPTerm "λ w ⇒ x" $ [< "w"] :\\ BVT 3,
|
||||||
parsesAs term fromPTerm' "λ x ⇒ x" $ [< "x"] :\\ BVT 0,
|
parsesAs term fromPTerm "λ x ⇒ x" $ [< "x"] :\\ BVT 0,
|
||||||
parsesAs term fromPTerm' "λ a b ⇒ f a b" $
|
parsesAs term fromPTerm "λ a b ⇒ f a b" $
|
||||||
[< "a", "b"] :\\ E (F "f" :@@ [BVT 1, BVT 0]),
|
[< "a", "b"] :\\ E (F "f" :@@ [BVT 1, BVT 0]),
|
||||||
parsesAs term fromPTerm' "f @𝑖" $
|
parsesAs term fromPTerm "f @𝑖" $
|
||||||
E $ F "f" :% BV 1
|
E $ F "f" :% BV 1
|
||||||
],
|
],
|
||||||
|
|
||||||
|
|
|
@ -5,6 +5,7 @@ import Quox.Syntax.Qty.Three
|
||||||
import Quox.Typechecker as Lib
|
import Quox.Typechecker as Lib
|
||||||
import public TypingImpls
|
import public TypingImpls
|
||||||
import TAP
|
import TAP
|
||||||
|
import Quox.EffExtra
|
||||||
|
|
||||||
|
|
||||||
data Error'
|
data Error'
|
||||||
|
@ -25,13 +26,10 @@ ToInfo Error' where
|
||||||
("wanted", prettyStr True bad)]
|
("wanted", prettyStr True bad)]
|
||||||
|
|
||||||
0 M : Type -> Type
|
0 M : Type -> Type
|
||||||
M = ReaderT (Definitions Three) $ Either Error'
|
M = Eff [Except Error', DefsReader Three]
|
||||||
|
|
||||||
inj : (forall m. CanTC Three m => m a) -> M a
|
inj : TC Three a -> M a
|
||||||
inj act = do
|
inj = rethrow . mapFst TCError <=< lift . runExcept
|
||||||
env <- ask
|
|
||||||
let res = runReaderT env act {m = Either (Typing.Error Three)}
|
|
||||||
either (throwError . TCError) pure res
|
|
||||||
|
|
||||||
|
|
||||||
reflTy : IsQty q => Term q d n
|
reflTy : IsQty q => Term q d n
|
||||||
|
@ -92,10 +90,12 @@ defGlobals = fromList
|
||||||
parameters (label : String) (act : Lazy (M ()))
|
parameters (label : String) (act : Lazy (M ()))
|
||||||
{default defGlobals globals : Definitions Three}
|
{default defGlobals globals : Definitions Three}
|
||||||
testTC : Test
|
testTC : Test
|
||||||
testTC = test label $ runReaderT globals act
|
testTC = test label {e = Error', a = ()} $
|
||||||
|
extract $ runExcept $ runReader globals act
|
||||||
|
|
||||||
testTCFail : Test
|
testTCFail : Test
|
||||||
testTCFail = testThrows label (const True) $ runReaderT globals act $> "()"
|
testTCFail = testThrows label (const True) $
|
||||||
|
(extract $ runExcept $ runReader globals act) $> "()"
|
||||||
|
|
||||||
|
|
||||||
anys : {n : Nat} -> QContext Three n
|
anys : {n : Nat} -> QContext Three n
|
||||||
|
@ -114,12 +114,10 @@ empty01 = eqDim (K Zero) (K One) empty
|
||||||
|
|
||||||
inferredTypeEq : TyContext Three d n -> (exp, got : Term Three d n) -> M ()
|
inferredTypeEq : TyContext Three d n -> (exp, got : Term Three d n) -> M ()
|
||||||
inferredTypeEq ctx exp got =
|
inferredTypeEq ctx exp got =
|
||||||
catchError
|
wrapErr (const $ WrongInfer exp got) $ inj $ equalType ctx exp got
|
||||||
(inj $ equalType ctx exp got)
|
|
||||||
(\_ : Error' => throwError $ WrongInfer exp got)
|
|
||||||
|
|
||||||
qoutEq : (exp, got : QOutput Three n) -> M ()
|
qoutEq : (exp, got : QOutput Three n) -> M ()
|
||||||
qoutEq qout res = unless (qout == res) $ throwError $ WrongQOut qout res
|
qoutEq qout res = unless (qout == res) $ throw $ WrongQOut qout res
|
||||||
|
|
||||||
inferAs : TyContext Three d n -> (sg : SQty Three) ->
|
inferAs : TyContext Three d n -> (sg : SQty Three) ->
|
||||||
Elim Three d n -> Term Three d n -> M ()
|
Elim Three d n -> Term Three d n -> M ()
|
||||||
|
@ -395,7 +393,7 @@ tests = "typechecker" :- [
|
||||||
testTC "0 · ℕ ⇐ ★₇" $ check_ empty szero Nat (TYPE 7),
|
testTC "0 · ℕ ⇐ ★₇" $ check_ empty szero Nat (TYPE 7),
|
||||||
testTCFail "1 · ℕ ⇍ ★₀" $ check_ empty sone Nat (TYPE 0),
|
testTCFail "1 · ℕ ⇍ ★₀" $ check_ empty sone Nat (TYPE 0),
|
||||||
testTC "1 · zero ⇐ ℕ" $ check_ empty sone Zero Nat,
|
testTC "1 · zero ⇐ ℕ" $ check_ empty sone Zero Nat,
|
||||||
testTC "1 · zero ⇍ ℕ×ℕ" $ check_ empty sone Zero (Nat `And` Nat),
|
testTCFail "1 · zero ⇍ ℕ×ℕ" $ check_ empty sone Zero (Nat `And` Nat),
|
||||||
testTC "ω·n : ℕ ⊢ 1 · succ n ⇐ ℕ" $
|
testTC "ω·n : ℕ ⊢ 1 · succ n ⇐ ℕ" $
|
||||||
check_ (ctx [< ("n", Nat)]) sone (Succ (BVT 0)) Nat,
|
check_ (ctx [< ("n", Nat)]) sone (Succ (BVT 0)) Nat,
|
||||||
testTC "1 · λ n ⇒ succ n ⇐ 1.ℕ → ℕ" $
|
testTC "1 · λ n ⇒ succ n ⇐ 1.ℕ → ℕ" $
|
||||||
|
|
|
@ -1,6 +1,6 @@
|
||||||
package quox-tests
|
package quox-tests
|
||||||
|
|
||||||
depends = base, contrib, elab-util, snocvect, quox-lib, tap
|
depends = base, contrib, elab-util, snocvect, quox-lib, tap, eff
|
||||||
|
|
||||||
executable = quox-tests
|
executable = quox-tests
|
||||||
main = Tests
|
main = Tests
|
||||||
|
|
Loading…
Reference in a new issue