quox/tests/Tests/FromPTerm.idr

113 lines
3.3 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
public export
data Failure =
ParseError Parser.Error
| FromParser FromParser.Error
| 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 FromParser.Error b)
(inp : String)
parameters {default (ltrim inp) label : String}
parsesWith : (b -> Bool) -> Test
parsesWith p = test label $ do
pres <- mapFst ParseError $ lexParseWith (grm "test") inp
res <- mapFst FromParser $ fromP pres
unless (p res) $ Left $ WrongResult $ show res
parses : Test
parses = parsesWith $ const True
%macro
parseMatch : TTImp -> Elab Test
parseMatch pat =
parsesWith <$> check `(\case ~(pat) => True; _ => False)
parsesAs : Eq b => b -> Test
parsesAs exp = parsesWith (== exp)
parameters {default "\{ltrim inp} # fails" label : String}
parseFails : Test
parseFails = test label $ do
pres <- mapFst ParseError $ lexParseWith (grm "test") inp
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
runFromParser : {default empty defs : Definitions} ->
Eff FromParserPure a -> Either FromParser.Error a
runFromParser = map fst . fst . fromParserPure 0 defs
export
tests : Test
tests = "PTerm → Term" :- [
"dimensions" :-
let fromPDim = runFromParser . 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 noLoc) (Zero noLoc) noLoc)]
-- 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 _ $ N $ E $ B (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) _) _)
],
todo "everything else"
]