more fromparser stuff
This commit is contained in:
parent
426c138c2b
commit
d9bc68446f
7 changed files with 130 additions and 31 deletions
|
@ -176,6 +176,7 @@ data FromPTermError =
|
||||||
| DuplicatesInEnum (List TagVal)
|
| DuplicatesInEnum (List TagVal)
|
||||||
| DimNotInScope Name
|
| DimNotInScope Name
|
||||||
| QtyNotGlobal PQty
|
| QtyNotGlobal PQty
|
||||||
|
| DimNameInTerm Name
|
||||||
|
|
||||||
public export
|
public export
|
||||||
FromPTerm : (Type -> Type) -> Type
|
FromPTerm : (Type -> Type) -> Type
|
||||||
|
@ -204,6 +205,11 @@ fromPDimWith : FromPTerm m =>
|
||||||
fromPDimWith ds (K e) = pure $ K e
|
fromPDimWith ds (K e) = pure $ K e
|
||||||
fromPDimWith ds (V i) = fromBaseName (pure . B) (throwError . DimNotInScope) ds i
|
fromPDimWith ds (V i) = fromBaseName (pure . B) (throwError . DimNotInScope) ds i
|
||||||
|
|
||||||
|
private
|
||||||
|
avoidDim : FromPTerm m => Context' BName d -> Name -> m (Term q d n)
|
||||||
|
avoidDim ds x =
|
||||||
|
fromName (const $ throwError $ DimNameInTerm x) (pure . FT) ds x
|
||||||
|
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
export
|
export
|
||||||
|
@ -263,7 +269,7 @@ mutual
|
||||||
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
|
map E $ (:%) <$> fromPTermElim ds ns s <*> fromPDimWith ds p
|
||||||
|
|
||||||
V x =>
|
V x =>
|
||||||
fromName (pure . E . B) (pure . FT) ns x
|
fromName (pure . E . B) (avoidDim ds) ns x
|
||||||
|
|
||||||
s :# a =>
|
s :# a =>
|
||||||
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a
|
map E $ (:#) <$> fromPTermWith ds ns s <*> fromPTermWith ds ns a
|
||||||
|
@ -315,9 +321,9 @@ fromPTerm = fromPTermWith [<] [<]
|
||||||
|
|
||||||
export
|
export
|
||||||
globalPQty : FromPTerm m => (q : PQty) -> m (IsGlobal q)
|
globalPQty : FromPTerm m => (q : PQty) -> m (IsGlobal q)
|
||||||
globalPQty Zero = pure GZero
|
globalPQty pi = case isGlobal pi of
|
||||||
globalPQty One = throwError $ QtyNotGlobal One
|
Yes y => pure y
|
||||||
globalPQty Any = pure GAny
|
No n => throwError $ QtyNotGlobal pi
|
||||||
|
|
||||||
|
|
||||||
-- [todo] extend substitutions so they can do this injection. that's the sort of
|
-- [todo] extend substitutions so they can do this injection. that's the sort of
|
||||||
|
|
|
@ -84,6 +84,12 @@ prettyDSubst th =
|
||||||
|
|
||||||
public export FromVar Dim where fromVar = B
|
public export FromVar Dim where fromVar = B
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
inject : Dim d -> Dim (d + inj)
|
||||||
|
inject (K e) = K e
|
||||||
|
inject (B i) = B $ inject i
|
||||||
|
|
||||||
export
|
export
|
||||||
CanShift Dim where
|
CanShift Dim where
|
||||||
K e // _ = K e
|
K e // _ = K e
|
||||||
|
@ -127,9 +133,3 @@ DecEq (Dim d) where
|
||||||
public export %inline
|
public export %inline
|
||||||
BV : (i : Nat) -> (0 _ : LT i d) => Dim d
|
BV : (i : Nat) -> (0 _ : LT i d) => Dim d
|
||||||
BV i = B $ V i
|
BV i = B $ V i
|
||||||
|
|
||||||
|
|
||||||
export
|
|
||||||
inject : {0 d' : Nat} -> Dim d -> Dim (d + d')
|
|
||||||
inject (K e) = K e
|
|
||||||
inject (B i) = B $ inject i
|
|
||||||
|
|
|
@ -77,7 +77,7 @@ fromNatWith (S i) (LTESucc p) = VS $ fromNatWith i p
|
||||||
|
|
||||||
public export %inline
|
public export %inline
|
||||||
V : (i : Nat) -> {auto 0 p : i `LT` n} -> Var n
|
V : (i : Nat) -> {auto 0 p : i `LT` n} -> Var n
|
||||||
V i {p} = fromNatWith i p
|
V i = fromNatWith i p
|
||||||
|
|
||||||
export %inline
|
export %inline
|
||||||
tryFromNat : Alternative f => (n : Nat) -> Nat -> f (Var n)
|
tryFromNat : Alternative f => (n : Nat) -> Nat -> f (Var n)
|
||||||
|
@ -112,7 +112,7 @@ toFromNat (S k) (LTESucc x) = cong S $ toFromNat k x
|
||||||
|
|
||||||
|
|
||||||
export
|
export
|
||||||
inject : Var m -> Var (m + n)
|
inject : Var m -> Var (m + inj)
|
||||||
inject VZ = VZ
|
inject VZ = VZ
|
||||||
inject (VS i) = VS $ inject i
|
inject (VS i) = VS $ inject i
|
||||||
|
|
||||||
|
|
|
@ -25,13 +25,14 @@ mutual
|
||||||
TYPE _ == _ = False
|
TYPE _ == _ = False
|
||||||
|
|
||||||
Pi qty1 arg1 res1 == Pi qty2 arg2 res2 =
|
Pi qty1 arg1 res1 == Pi qty2 arg2 res2 =
|
||||||
qty1 == qty2 && arg1 == arg2 && res1 == res2
|
qty1 == qty2 && arg1 == arg2 && res1.term == res2.term
|
||||||
Pi {} == _ = False
|
Pi {} == _ = False
|
||||||
|
|
||||||
Lam body1 == Lam body2 = body1 == body2
|
Lam body1 == Lam body2 = body1.term == body2.term
|
||||||
Lam {} == _ = False
|
Lam {} == _ = False
|
||||||
|
|
||||||
Sig fst1 snd1 == Sig fst2 snd2 = fst1 == fst2 && snd1 == snd2
|
Sig fst1 snd1 == Sig fst2 snd2 =
|
||||||
|
fst1 == fst2 && snd1.term == snd2.term
|
||||||
Sig {} == _ = False
|
Sig {} == _ = False
|
||||||
|
|
||||||
Pair fst1 snd1 == Pair fst2 snd2 = fst1 == fst2 && snd1 == snd2
|
Pair fst1 snd1 == Pair fst2 snd2 = fst1 == fst2 && snd1 == snd2
|
||||||
|
@ -44,10 +45,10 @@ mutual
|
||||||
Tag _ == _ = False
|
Tag _ == _ = False
|
||||||
|
|
||||||
Eq ty1 l1 r1 == Eq ty2 l2 r2 =
|
Eq ty1 l1 r1 == Eq ty2 l2 r2 =
|
||||||
ty1 == ty2 && l1 == l2 && r1 == r2
|
ty1.term == ty2.term && l1 == l2 && r1 == r2
|
||||||
Eq {} == _ = False
|
Eq {} == _ = False
|
||||||
|
|
||||||
DLam body1 == DLam body2 = body1 == body2
|
DLam body1 == DLam body2 = body1.term == body2.term
|
||||||
DLam {} == _ = False
|
DLam {} == _ = False
|
||||||
|
|
||||||
E e == E f = e == f
|
E e == E f = e == f
|
||||||
|
@ -77,11 +78,11 @@ mutual
|
||||||
(_ :@ _) == _ = False
|
(_ :@ _) == _ = False
|
||||||
|
|
||||||
CasePair q1 p1 r1 b1 == CasePair q2 p2 r2 b2 =
|
CasePair q1 p1 r1 b1 == CasePair q2 p2 r2 b2 =
|
||||||
q1 == q2 && p1 == p2 && r1 == r2 && b1 == b2
|
q1 == q2 && p1 == p2 && r1.term == r2.term && b1.term == b2.term
|
||||||
CasePair {} == _ = False
|
CasePair {} == _ = False
|
||||||
|
|
||||||
CaseEnum q1 t1 r1 a1 == CaseEnum q2 t2 r2 a2 =
|
CaseEnum q1 t1 r1 a1 == CaseEnum q2 t2 r2 a2 =
|
||||||
q1 == q2 && t1 == t2 && r1 == r2 && a1 == a2
|
q1 == q2 && t1 == t2 && r1.term == r2.term && a1 == a2
|
||||||
CaseEnum {} == _ = False
|
CaseEnum {} == _ = False
|
||||||
|
|
||||||
(fun1 :% dim1) == (fun2 :% dim2) = fun1 == fun2 && dim1 == dim2
|
(fun1 :% dim1) == (fun2 :% dim2) = fun1 == fun2 && dim1 == dim2
|
||||||
|
@ -102,16 +103,6 @@ mutual
|
||||||
Nothing => False
|
Nothing => False
|
||||||
DCloE {} == _ = False
|
DCloE {} == _ = False
|
||||||
|
|
||||||
export covering
|
|
||||||
(forall n. Eq (f n)) => Eq (ScopedBody s f n) where
|
|
||||||
Y x == Y y = x == y
|
|
||||||
N x == N y = x == y
|
|
||||||
_ == _ = False
|
|
||||||
|
|
||||||
export covering
|
|
||||||
(forall n. Eq (f n)) => Eq (Scoped s f n) where
|
|
||||||
S _ x == S _ y = x == y
|
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
PrettyHL q => Show (Term q d n) where
|
PrettyHL q => Show (Term q d n) where
|
||||||
showPrec d t = showParens (d /= Open) $ prettyStr True t
|
showPrec d t = showParens (d /= Open) $ prettyStr True t
|
||||||
|
|
|
@ -6,6 +6,7 @@ import Tests.Equal
|
||||||
import Tests.Typechecker
|
import Tests.Typechecker
|
||||||
import Tests.Lexer
|
import Tests.Lexer
|
||||||
import Tests.Parser
|
import Tests.Parser
|
||||||
|
import Tests.FromPTerm
|
||||||
import System
|
import System
|
||||||
|
|
||||||
|
|
||||||
|
@ -15,7 +16,8 @@ allTests = [
|
||||||
Equal.tests,
|
Equal.tests,
|
||||||
Typechecker.tests,
|
Typechecker.tests,
|
||||||
Lexer.tests,
|
Lexer.tests,
|
||||||
Parser.tests
|
Parser.tests,
|
||||||
|
FromPTerm.tests
|
||||||
]
|
]
|
||||||
|
|
||||||
main : IO ()
|
main : IO ()
|
||||||
|
|
93
tests/Tests/FromPTerm.idr
Normal file
93
tests/Tests/FromPTerm.idr
Normal file
|
@ -0,0 +1,93 @@
|
||||||
|
module Tests.FromPTerm
|
||||||
|
|
||||||
|
import Quox.Parser.Syntax
|
||||||
|
import Quox.Parser
|
||||||
|
import TermImpls
|
||||||
|
import TAP
|
||||||
|
|
||||||
|
import Derive.Prelude
|
||||||
|
%language ElabReflection
|
||||||
|
|
||||||
|
public export
|
||||||
|
data Failure =
|
||||||
|
LexError Lexer.Error
|
||||||
|
| ParseError (List1 (ParsingError Token))
|
||||||
|
| FromPTermError FromPTermError
|
||||||
|
| WrongResult String
|
||||||
|
| ExpectedFail String
|
||||||
|
|
||||||
|
%runElab derive "Syntax.FromPTermError" [Show]
|
||||||
|
|
||||||
|
export
|
||||||
|
ToInfo Failure where
|
||||||
|
toInfo (LexError err) =
|
||||||
|
[("type", "LexError"), ("info", show err)]
|
||||||
|
toInfo (ParseError errs) =
|
||||||
|
("type", "ParseError") ::
|
||||||
|
map (bimap show show) ([1 .. length errs] `zip` toList errs)
|
||||||
|
toInfo (FromPTermError err) =
|
||||||
|
[("type", "FromPTermError"),
|
||||||
|
("got", show err)]
|
||||||
|
toInfo (WrongResult got) =
|
||||||
|
[("type", "WrongResult"), ("got", got)]
|
||||||
|
toInfo (ExpectedFail got) =
|
||||||
|
[("type", "ExpectedFail"), ("got", got)]
|
||||||
|
|
||||||
|
|
||||||
|
parameters {c : Bool} {auto _ : Show b}
|
||||||
|
(grm : Grammar c a) (fromP : a -> Either FromPTermError b)
|
||||||
|
(inp : String)
|
||||||
|
parameters {default (ltrim inp) label : String}
|
||||||
|
parsesWith : (b -> Bool) -> Test
|
||||||
|
parsesWith p = test label $ do
|
||||||
|
toks <- mapFst LexError $ lex inp
|
||||||
|
(pres, _) <- mapFst ParseError $ parse (grm <* eof) toks
|
||||||
|
res <- mapFst FromPTermError $ fromP pres
|
||||||
|
unless (p res) $ Left $ WrongResult $ show res
|
||||||
|
|
||||||
|
parses : Test
|
||||||
|
parses = parsesWith $ const True
|
||||||
|
|
||||||
|
parsesAs : Eq b => b -> Test
|
||||||
|
parsesAs exp = parsesWith (== exp)
|
||||||
|
|
||||||
|
parameters {default "\{ltrim inp} # fails" label : String}
|
||||||
|
parseFails : Test
|
||||||
|
parseFails = test label $ do
|
||||||
|
toks <- mapFst LexError $ lex inp
|
||||||
|
(pres, _) <- mapFst ParseError $ parse (grm <* eof) toks
|
||||||
|
either (const $ Right ()) (Left . ExpectedFail . show) $ fromP pres
|
||||||
|
|
||||||
|
|
||||||
|
FromString BName where fromString = Just . fromString
|
||||||
|
|
||||||
|
export
|
||||||
|
tests : Test
|
||||||
|
tests = "PTerm → Term" :- [
|
||||||
|
"dimensions" :-
|
||||||
|
let fromPDim = fromPDimWith [< "𝑖", "𝑗"]
|
||||||
|
in [
|
||||||
|
note "dim ctx: [𝑖, 𝑗]",
|
||||||
|
parsesAs dim fromPDim "𝑖" (BV 1),
|
||||||
|
parsesAs dim fromPDim "𝑗" (BV 0),
|
||||||
|
parseFails dim fromPDim "𝑘",
|
||||||
|
parsesAs dim fromPDim "0" (K Zero),
|
||||||
|
parsesAs dim fromPDim "1" (K One)
|
||||||
|
],
|
||||||
|
|
||||||
|
"terms" :-
|
||||||
|
let fromPTerm' = fromPTermWith [< "𝑖", "𝑗"] [< "A", "x", "y", "z"]
|
||||||
|
in [
|
||||||
|
note "dim ctx: [𝑖, 𝑗]; term ctx: [A, x, y, z]",
|
||||||
|
parsesAs term fromPTerm' "x" $ BVT 2,
|
||||||
|
parseFails term fromPTerm' "𝑖",
|
||||||
|
parsesAs term fromPTerm' "f" $ FT "f",
|
||||||
|
parsesAs term fromPTerm' "λ w ⇒ w" $ ["w"] :\\ BVT 0,
|
||||||
|
parsesAs term fromPTerm' "λ w ⇒ x" $ ["w"] :\\ BVT 3,
|
||||||
|
parsesAs term fromPTerm' "λ x ⇒ x" $ ["x"] :\\ BVT 0,
|
||||||
|
parsesAs term fromPTerm' "λ a b ⇒ f a b" $
|
||||||
|
["a", "b"] :\\ E (F "f" :@@ [BVT 1, BVT 0]),
|
||||||
|
parsesAs term fromPTerm' "f @𝑖" $
|
||||||
|
E $ F "f" :% BV 1
|
||||||
|
]
|
||||||
|
]
|
|
@ -396,6 +396,13 @@ tests = "typechecker" :- [
|
||||||
(Pi_ One "x" (FT "A")
|
(Pi_ One "x" (FT "A")
|
||||||
(Eq0 (E $ F "P" :@ BVT 0)
|
(Eq0 (E $ F "P" :@ BVT 0)
|
||||||
(E $ F "p" :@ BVT 0) (E $ F "q" :@ BVT 0)))
|
(E $ F "p" :@ BVT 0) (E $ F "q" :@ BVT 0)))
|
||||||
(Eq0 (Pi_ Any "x" (FT "A") $ E $ F "P" :@ BVT 0) (FT "p") (FT "q")))
|
(Eq0 (Pi_ Any "x" (FT "A") $ E $ F "P" :@ BVT 0) (FT "p") (FT "q"))),
|
||||||
|
todo "absurd (when coerce is in)"
|
||||||
|
-- absurd : (`true ≡ `false : {true, false}) ⇾ (0·A : ★₀) → A ≔
|
||||||
|
-- λ e ⇒
|
||||||
|
-- case coerce [i ⇒ case e @i return ★₀ of {`true ⇒ {tt}; `false ⇒ {}}]
|
||||||
|
-- @0 @1 `tt
|
||||||
|
-- return A
|
||||||
|
-- of { }
|
||||||
]
|
]
|
||||||
]
|
]
|
||||||
|
|
Loading…
Reference in a new issue