quox/tests/Tests/FromPTerm.idr

108 lines
3.2 KiB
Idris
Raw Blame History

This file contains ambiguous Unicode characters

This file contains Unicode characters that might be confused with other characters. If you think that this is intentional, you can safely ignore this warning. Use the Escape button to reveal them.

module Tests.FromPTerm
import Quox.Parser.FromParser
import Quox.Parser
import TypingImpls
import Tests.Parser as TParser
import Quox.EffExtra
import TAP
import AstExtra
import PrettyExtra
import System.File
import Derive.Prelude
%language ElabReflection
%hide TParser.Failure
%hide TParser.ExpectedFail
PError = Parser.Error
FPError = FromParser.Error
public export
data Failure =
ParseError PError
| FromParser FPError
| WrongResult String
| ExpectedFail String
%runElab derive "FileError" [Show]
export
ToInfo Failure where
toInfo (ParseError err) = toInfo err
toInfo (FromParser err) =
[("type", "FromParserError"),
("got", prettyStr $ prettyError True err)]
toInfo (WrongResult got) =
[("type", "WrongResult"), ("got", got)]
toInfo (ExpectedFail got) =
[("type", "ExpectedFail"), ("got", got)]
parameters {c : Bool} {auto _ : Show b}
(grm : FileName -> Grammar c a)
(fromP : a -> Either FPError b)
(inp : String)
parsesWith : String -> (b -> Bool) -> Test
parsesWith label p = test label $ do
pres <- mapFst ParseError $ lexParseWith (grm "test") inp
res <- mapFst FromParser $ fromP pres
unless (p res) $ Left $ WrongResult $ show res
%macro
parseMatch : {default (ltrim inp) label : String} -> TTImp -> Elab Test
parseMatch {label} pat =
parsesWith label <$> check `(\case ~(pat) => True; _ => False)
parseFails : {default "\{ltrim inp} # fails" label : String} -> Test
parseFails {label} = test label $ do
pres <- mapFst ParseError $ lexParseWith (grm "test") inp
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
runFromParser : Definitions -> Eff FromParserPure a -> Either FPError a
runFromParser defs = map val . fromParserPure [<] 0 defs initStack
export
tests : Test
tests = "PTerm → Term" :- [
"dimensions" :-
let fromPDim = runFromParser empty . fromPDimWith [< "𝑖", "𝑗"]
in [
note "dim ctx: [𝑖, 𝑗]",
parseMatch dim fromPDim "𝑖" `(B (VS VZ) _),
parseMatch dim fromPDim "𝑗" `(B VZ _),
parseFails dim fromPDim "𝑘",
parseMatch dim fromPDim "0" `(K Zero _),
parseMatch dim fromPDim "1" `(K One _)
],
"terms" :-
let defs = fromList [("f", mkDef GAny (^NAT) (^Zero))]
-- doesn't have to be well typed yet, just well scoped
fromPTerm = runFromParser defs .
fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
in [
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",
parseMatch term fromPTerm "x" `(E $ B (VS $ VS VZ) _),
parseFails term fromPTerm "𝑖",
parseMatch term fromPTerm "f" `(E $ F "f" {}),
parseMatch term fromPTerm "λ w ⇒ w"
`(Lam (S _ $ Y $ E $ B VZ _) _),
parseMatch term fromPTerm "λ w ⇒ x"
`(Lam (S _ $ Y $ E $ B (VS $ VS $ VS VZ) _) _),
parseMatch term fromPTerm "λ x ⇒ x"
`(Lam (S _ $ Y $ E $ B VZ _) _),
parseMatch term fromPTerm "λ a b ⇒ f a b"
`(Lam (S _ $ Y $
Lam (S _ $ Y $
E $ App (App (F "f" {}) (E $ B (VS VZ) _) _) (E $ B VZ _) _) _) _),
parseMatch term fromPTerm "f @𝑖" $
`(E $ DApp (F "f" {}) (B (VS VZ) _) _),
parseFails term fromPTerm "λ x ⇒ x¹"
],
todo "everything else"
]