remove src directories
This commit is contained in:
parent
79211cff84
commit
804f1e3638
36 changed files with 0 additions and 3 deletions
101
tests/Tests/Equal.idr
Normal file
101
tests/Tests/Equal.idr
Normal file
|
@ -0,0 +1,101 @@
|
|||
module Tests.Equal
|
||||
|
||||
import Quox.Equal
|
||||
import Quox.Pretty
|
||||
import TAP
|
||||
|
||||
export
|
||||
ToInfo Equal.Error where
|
||||
toInfo (ClashT mode s t) =
|
||||
[("clash", "term"),
|
||||
("mode", show mode),
|
||||
("left", prettyStr True s),
|
||||
("right", prettyStr True t)]
|
||||
toInfo (ClashU mode k l) =
|
||||
[("clash", "universe"),
|
||||
("mode", show mode),
|
||||
("left", prettyStr True k),
|
||||
("right", prettyStr True l)]
|
||||
toInfo (ClashQ pi rh) =
|
||||
[("clash", "quantity"),
|
||||
("left", prettyStr True pi),
|
||||
("right", prettyStr True rh)]
|
||||
|
||||
|
||||
M = Either Equal.Error
|
||||
|
||||
testEq : String -> Lazy (M ()) -> Test
|
||||
testEq = test
|
||||
|
||||
testNeq : String -> Lazy (M ()) -> Test
|
||||
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
|
||||
|
||||
equalT : {default 0 d, n : Nat} -> Term d n -> Term d n -> M ()
|
||||
equalT = Quox.Equal.equalT
|
||||
%hide Quox.Equal.equalT
|
||||
|
||||
subE : {default 0 d, n : Nat} -> Elim d n -> Elim d n -> M ()
|
||||
subE = Quox.Equal.subE
|
||||
%hide Quox.Equal.subE
|
||||
|
||||
equalE : {default 0 d, n : Nat} -> Elim d n -> Elim d n -> M ()
|
||||
equalE = Quox.Equal.equalE
|
||||
%hide Quox.Equal.equalE
|
||||
|
||||
|
||||
export
|
||||
tests : Test
|
||||
tests = "equality & subtyping" :- [
|
||||
"universes" :- [
|
||||
testEq "Type 0 == Type 0" $
|
||||
equalT (TYPE (U 0)) (TYPE (U 0)),
|
||||
testNeq "Type 0 =/= Type 1" $
|
||||
equalT (TYPE (U 0)) (TYPE (U 1)),
|
||||
testNeq "Type 1 =/= Type 0" $
|
||||
equalT (TYPE (U 1)) (TYPE (U 0)),
|
||||
testEq "Type 0 <: Type 0" $
|
||||
subT (TYPE (U 0)) (TYPE (U 0)),
|
||||
testEq "Type 0 <: Type 1" $
|
||||
subT (TYPE (U 0)) (TYPE (U 1)),
|
||||
testNeq "Type 1 </: Type 0" $
|
||||
subT (TYPE (U 1)) (TYPE (U 0))
|
||||
],
|
||||
todo "pi",
|
||||
todo "lambda",
|
||||
todo "term closure",
|
||||
todo "term d-closure",
|
||||
|
||||
"free var" :- [
|
||||
testEq "A == A" $
|
||||
equalE (F "A") (F "A"),
|
||||
testNeq "A =/= B" $
|
||||
equalE (F "A") (F "B"),
|
||||
testEq "A <: A" $
|
||||
subE (F "A") (F "A"),
|
||||
testNeq "A </: B" $
|
||||
subE (F "A") (F "B")
|
||||
],
|
||||
todo "bound var",
|
||||
"application" :-
|
||||
let a = F "a"; a' = E a
|
||||
A = FT "A"
|
||||
λxx = Lam "x" (TUsed (BVT 0))
|
||||
A_A = Arr one A A
|
||||
λxx' = λxx :# A_A
|
||||
in [
|
||||
testEq "(λx. x : A -> A) a == ((a : A) : A) (β)" $
|
||||
equalE (λxx' :@ a') (E (a' :# A) :# A),
|
||||
testEq "(λx. x : _) a == a (βυ)" $
|
||||
equalE (λxx' :@ a') a
|
||||
],
|
||||
todo "annotation",
|
||||
todo "elim closure",
|
||||
todo "elim d-closure",
|
||||
|
||||
todo "clashes"
|
||||
]
|
154
tests/Tests/Lexer.idr
Normal file
154
tests/Tests/Lexer.idr
Normal file
|
@ -0,0 +1,154 @@
|
|||
module Tests.Lexer
|
||||
|
||||
import Quox.Lexer
|
||||
import TAP
|
||||
|
||||
|
||||
RealError = Quox.Lexer.Error
|
||||
%hide Quox.Lexer.Error
|
||||
|
||||
export
|
||||
ToInfo RealError where
|
||||
toInfo (Err reason line col char) =
|
||||
[("reason", show reason),
|
||||
("line", show line),
|
||||
("col", show col),
|
||||
("char", show char)]
|
||||
|
||||
data Error
|
||||
= LexerError RealError
|
||||
| WrongAnswer (List Token) (List Token)
|
||||
| TestFailed (List Token)
|
||||
|
||||
ToInfo Error where
|
||||
toInfo (LexerError err) = toInfo err
|
||||
toInfo (WrongAnswer exp got) =
|
||||
[("expected", show exp), ("received", show got)]
|
||||
toInfo (TestFailed got) =
|
||||
[("failed", show got)]
|
||||
|
||||
|
||||
lex' : String -> Either Error (List Token)
|
||||
lex' = bimap LexerError (map val) . lex
|
||||
|
||||
parameters (label : String) (input : String)
|
||||
acceptsSuchThat' : (List Token -> Maybe Error) -> Test
|
||||
acceptsSuchThat' p = test label $ delay $ do
|
||||
res <- bimap LexerError (map val) $ lex input
|
||||
case p res of
|
||||
Just err => throwError err
|
||||
Nothing => pure ()
|
||||
|
||||
acceptsSuchThat : (List Token -> Bool) -> Test
|
||||
acceptsSuchThat p = acceptsSuchThat' $ \res =>
|
||||
if p res then Nothing else Just $ TestFailed res
|
||||
|
||||
acceptsWith : List Token -> Test
|
||||
acceptsWith expect = acceptsSuchThat' $ \res =>
|
||||
if res == expect then Nothing else Just $ WrongAnswer expect res
|
||||
|
||||
accepts : Test
|
||||
accepts = acceptsSuchThat $ const True
|
||||
|
||||
rejects : Test
|
||||
rejects = testThrows label (\case LexerError _ => True; _ => False) $ delay $
|
||||
bimap LexerError (map val) $ lex input
|
||||
|
||||
parameters (input : String) {default False esc : Bool}
|
||||
show' : String -> String
|
||||
show' s = if esc then show s else "\"\{s}\""
|
||||
|
||||
acceptsWith' : List Token -> Test
|
||||
acceptsWith' = acceptsWith (show' input) input
|
||||
|
||||
accepts' : Test
|
||||
accepts' = accepts (show' input) input
|
||||
|
||||
rejects' : Test
|
||||
rejects' = rejects "\{show' input} (reject)" input
|
||||
|
||||
|
||||
tests = "lexer" :- [
|
||||
"comments" :- [
|
||||
acceptsWith' "" [],
|
||||
acceptsWith' " \n \t\t " [] {esc = True},
|
||||
acceptsWith' "-- a" [],
|
||||
acceptsWith' "{- -}" [],
|
||||
acceptsWith' "{--}" [],
|
||||
acceptsWith' "{------}" [],
|
||||
acceptsWith' "{- {- -} -}" [],
|
||||
acceptsWith' "{- } -}" [],
|
||||
rejects' "{-}",
|
||||
rejects' "{- {- -}",
|
||||
acceptsWith' "( -- comment \n )" [P LParen, P RParen] {esc = True}
|
||||
],
|
||||
|
||||
"punctuation" :- [
|
||||
acceptsWith' "({[:,::]})"
|
||||
[P LParen, P LBrace, P LSquare,
|
||||
P Colon, P Comma, P DblColon,
|
||||
P RSquare, P RBrace, P RParen],
|
||||
acceptsWith' " ( { [ : , :: ] } ) "
|
||||
[P LParen, P LBrace, P LSquare,
|
||||
P Colon, P Comma, P DblColon,
|
||||
P RSquare, P RBrace, P RParen],
|
||||
acceptsWith' "-> → => ⇒ ** × << ⊲ ∷"
|
||||
[P Arrow, P Arrow, P DblArrow, P DblArrow,
|
||||
P Times, P Times, P Triangle, P Triangle, P DblColon],
|
||||
acceptsWith' "_" [P Wild]
|
||||
],
|
||||
|
||||
"names & symbols" :- [
|
||||
acceptsWith' "a" [Name "a"],
|
||||
acceptsWith' "abc" [Name "abc"],
|
||||
acceptsWith' "_a" [Name "_a"],
|
||||
acceptsWith' "a_" [Name "a_"],
|
||||
acceptsWith' "a_b" [Name "a_b"],
|
||||
acceptsWith' "abc'" [Name "abc'"],
|
||||
acceptsWith' "a'b'c''" [Name "a'b'c''"],
|
||||
acceptsWith' "abc123" [Name "abc123"],
|
||||
acceptsWith' "_1" [Name "_1"],
|
||||
acceptsWith' "ab cd" [Name "ab", Name "cd"],
|
||||
acceptsWith' "ab{--}cd" [Name "ab", Name "cd"],
|
||||
acceptsWith' "'a" [Symbol "a"],
|
||||
acceptsWith' "'ab" [Symbol "ab"],
|
||||
acceptsWith' "'_b" [Symbol "_b"],
|
||||
acceptsWith' "a.b.c" [Name "a", P Dot, Name "b", P Dot, Name "c"],
|
||||
rejects' "'",
|
||||
rejects' "1abc"
|
||||
],
|
||||
|
||||
"keywords" :- [
|
||||
acceptsWith' "fun" [K Fun],
|
||||
acceptsWith' "λ" [K Fun],
|
||||
acceptsWith' "let" [K Let],
|
||||
acceptsWith' "in" [K In],
|
||||
acceptsWith' "case" [K Case],
|
||||
acceptsWith' "of" [K Of],
|
||||
acceptsWith' "ω" [K Omega],
|
||||
acceptsWith' "#" [K Omega],
|
||||
acceptsWith' "funk" [Name "funk"]
|
||||
],
|
||||
|
||||
"numbers" :- [
|
||||
acceptsWith' "0" [N Zero],
|
||||
acceptsWith' "1" [N One],
|
||||
acceptsWith' "2" [N $ Other 2],
|
||||
acceptsWith' "69" [N $ Other 69],
|
||||
acceptsWith' "1_000" [N $ Other 1000],
|
||||
acceptsWith' "0o0" [N Zero],
|
||||
acceptsWith' "0o105" [N $ Other 69],
|
||||
acceptsWith' "0O0" [N Zero],
|
||||
acceptsWith' "0O105" [N $ Other 69],
|
||||
acceptsWith' "0x0" [N Zero],
|
||||
acceptsWith' "0x45" [N $ Other 69],
|
||||
acceptsWith' "0xabc" [N $ Other 2748],
|
||||
acceptsWith' "0xABC" [N $ Other 2748],
|
||||
acceptsWith' "0xA_BC" [N $ Other 2748],
|
||||
acceptsWith' "0X0" [N Zero],
|
||||
acceptsWith' "0X45" [N $ Other 69],
|
||||
acceptsWith' "0XABC" [N $ Other 2748],
|
||||
rejects' "1_",
|
||||
rejects' "1__000"
|
||||
]
|
||||
]
|
77
tests/Tests/Parser.idr
Normal file
77
tests/Tests/Parser.idr
Normal file
|
@ -0,0 +1,77 @@
|
|||
module Tests.Parser
|
||||
|
||||
import Quox.Syntax
|
||||
import Quox.Parser
|
||||
import Quox.Lexer
|
||||
import Tests.Lexer
|
||||
|
||||
import Data.SnocVect
|
||||
import Text.Parser
|
||||
import TAP
|
||||
|
||||
|
||||
export
|
||||
Show tok => ToInfo (ParsingError tok) where
|
||||
toInfo (Error msg Nothing) = [("msg", msg)]
|
||||
toInfo (Error msg (Just loc)) = [("loc", show loc), ("msg", msg)]
|
||||
|
||||
|
||||
numberErrs : List1 Info -> Info
|
||||
numberErrs (head ::: []) = head
|
||||
numberErrs (head ::: tail) = go 0 (head :: tail) where
|
||||
number1 : Nat -> Info -> Info
|
||||
number1 n = map $ \(k, v) => (show n ++ k, v)
|
||||
|
||||
go : Nat -> List Info -> Info
|
||||
go k [] = []
|
||||
go k (x :: xs) = number1 k x ++ go (S k) xs
|
||||
|
||||
export
|
||||
ToInfo Parser.Error where
|
||||
toInfo (Lex err) = toInfo err
|
||||
toInfo (Parse errs) = numberErrs $ map toInfo errs
|
||||
toInfo (Leftover toks) = toInfo [("leftover", toks)]
|
||||
|
||||
|
||||
RealError = Quox.Parser.Error
|
||||
%hide Lexer.RealError
|
||||
%hide Quox.Parser.Error
|
||||
|
||||
data Error a
|
||||
= Parser RealError
|
||||
| Unexpected a a
|
||||
| ShouldFail a
|
||||
|
||||
export
|
||||
Show a => ToInfo (Error a) where
|
||||
toInfo (Parser err) = toInfo err
|
||||
toInfo (Unexpected exp got) = toInfo $
|
||||
[("expected", exp), ("received", got)]
|
||||
toInfo (ShouldFail got) = toInfo [("success", got)]
|
||||
|
||||
|
||||
parameters {c : Bool} (grm : Grammar c a) (input : String)
|
||||
parses : (Show a, Eq a) => a -> Test
|
||||
parses exp = test input $ delay $
|
||||
case lexParseAll grm input of
|
||||
Right got => if got == exp then Right ()
|
||||
else Left $ Unexpected exp got
|
||||
Left err => Left $ Parser err
|
||||
|
||||
rejects : Show a => Test
|
||||
rejects = test input $ do
|
||||
case lexParseAll grm input of
|
||||
Left err => Right ()
|
||||
Right val => Left $ ShouldFail val
|
||||
|
||||
tests = "parser" :- [
|
||||
"numbers" :- let parses = parses number in [
|
||||
parses "0" 0,
|
||||
parses "1" 1,
|
||||
parses "1000" 1000
|
||||
],
|
||||
|
||||
"bound vars" :- let parses = parses (bound [< "x", "y", "z"]) in [
|
||||
|
||||
]
|
||||
]
|
Loading…
Add table
Add a link
Reference in a new issue