add definitions to parser

This commit is contained in:
rhiannon morris 2023-03-06 11:35:57 +01:00
parent ab2508e0ce
commit 757ea89b0f
9 changed files with 199 additions and 29 deletions

View File

@ -215,7 +215,11 @@ reserved =
W "λ" `Or` W "fun",
W "δ" `Or` W "dfun",
W "ω" `Or` S "#",
S "" `Or` W "Type"]
S "" `Or` W "Type",
W1 "def",
W1 "def0",
W "defω" `Or` W "def#",
S "" `Or` S ":="]
||| `IsReserved str` is true if `R str` might actually show up in
||| the token stream

View File

@ -21,18 +21,6 @@ public export
Grammar = Core.Grammar () Token
%hide Core.Grammar
private
data PArg = T PTerm | D PDim
private
apply : PTerm -> List PArg -> PTerm
apply = foldl $ \f, x => case x of T x => f :@ x; D p => f :% p
private
annotate : PTerm -> Maybe PTerm -> PTerm
annotate s a = maybe s (s :#) a
export
res : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
@ -167,19 +155,15 @@ private
makeSig : MakeBinder 0
makeSig = ("×", \([], x, s) => Sig x s)
private
plural : Nat -> a -> a -> a
plural 1 s p = s
plural _ s p = p
private
makeBinder : {m, n : Nat} -> MakeBinder m -> PBinderHead n -> PTerm ->
Grammar False PTerm
makeBinder (str, f) h t =
case decEq m n of
Yes Refl => pure $ f h t
No _ => fatalError
"'\{str}' expects \{show m} quantit\{plural m "y" "ies"}, got \{show n}"
No _ =>
let q = if m == 1 then "quantity" else "quantities" in
fatalError "'\{str}' expects \{show m} \{q}, got \{show n}"
private
binderInfix : Grammar True (n ** MakeBinder n)
@ -208,6 +192,9 @@ mutual
<|> caseTerm
<|> bindTerm
<|> [|annotate infixEqTerm (optional $ resC "" *> term)|]
where
annotate : PTerm -> Maybe PTerm -> PTerm
annotate s a = maybe s (s :#) a
private covering
lamTerm : Grammar True PTerm
@ -248,11 +235,15 @@ mutual
appTerm = resC "" *> [|TYPE nat|]
<|> resC "Eq" *> [|Eq (bracks optBinderTerm) aTerm aTerm|]
<|> [|apply aTerm (many appArg)|]
where
data PArg = T PTerm | D PDim
private covering
appArg : Grammar True PArg
appArg = [|D $ resC "@" *> dim|]
<|> [|T aTerm|]
appArg : Grammar True PArg
appArg = [|D $ resC "@" *> dim|]
<|> [|T aTerm|]
apply : PTerm -> List PArg -> PTerm
apply = foldl $ \f, x => case x of T x => f :@ x; D p => f :% p
private covering
aTerm : Grammar True PTerm
@ -269,8 +260,24 @@ mutual
name <- bname
resC ":"
ty <- term
pure (_ ** (qs.snd, name, ty))
pure (qs.fst ** (qs.snd, name, ty))
private covering
optBinderTerm : Grammar True (BName, PTerm)
optBinderTerm = [|MkPair optNameBinder term|]
export covering
defIntro : Grammar True PQty
defIntro = Zero <$ resC "def0"
<|> Any <$ resC "defω"
<|> resC "def" *> option Any (qty <* resC "·")
export covering
definition : Grammar True PDefinition
definition =
[|MkPDef defIntro name (resC ":" *> term) (resC "" *> term)|] <* resC ";"
export covering
input : Grammar False (List PDefinition)
input = many definition

View File

@ -2,6 +2,7 @@ module Quox.Parser.Syntax
import public Quox.Syntax
import public Quox.Syntax.Qty.Three
import public Quox.Definition
import public Control.Monad.Either
@ -65,6 +66,17 @@ namespace PTerm
%runElab deriveMutual ["PTerm", "PCaseBody"] [Eq, Ord, Show]
public export
record PDefinition where
constructor MkPDef
qty : PQty
name : Name
type : PTerm
term : PTerm
%name PDefinition def
%runElab derive "PDefinition" [Eq, Ord, Show]
export
toPDimWith : Context' BaseName d -> Dim d -> PDim
toPDimWith ds (K e) = K e
@ -163,6 +175,7 @@ data FromPTermError =
AnnotationNeeded PTerm
| DuplicatesInEnum (List TagVal)
| DimNotInScope Name
| QtyNotGlobal PQty
public export
FromPTerm : (Type -> Type) -> Type
@ -189,8 +202,8 @@ export
fromPDimWith : FromPTerm m =>
Context' BName d -> PDim -> m (Dim d)
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
mutual
export
@ -298,3 +311,22 @@ mutual
export %inline
fromPTerm : FromPTerm m => PTerm -> m (Term Three 0 0)
fromPTerm = fromPTermWith [<] [<]
export
globalPQty : FromPTerm m => (q : PQty) -> m (IsGlobal q)
globalPQty Zero = pure GZero
globalPQty One = throwError $ QtyNotGlobal One
globalPQty Any = pure GAny
-- [todo] extend substitutions so they can do this injection. that's the sort of
-- thing they are for.
export
fromPDefinition : FromPTerm m => PDefinition -> m (Name, Definition Three)
fromPDefinition (MkPDef {name, qty, type, term}) =
pure (name, MkDef' {
qty, qtyGlobal = !(globalPQty qty),
type = T $ inject !(fromPTerm type),
term = Just $ T $ inject !(fromPTerm term)
})

View File

@ -127,3 +127,9 @@ DecEq (Dim d) where
public export %inline
BV : (i : Nat) -> (0 _ : LT i d) => Dim d
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

View File

@ -120,6 +120,21 @@ ssDownViaNat by =
%transform "Shift.ssDown" ssDown = ssDownViaNat
public export
inject : Shift from to -> Shift (from + n') (to + n')
inject SZ = SZ
inject (SS by) = SS (inject by)
public export
injectViaNat : Shift from to -> Shift (from + n') (to + n')
injectViaNat by =
rewrite shiftDiff by in
rewrite sym $ plusAssociative by.nat from n' in
fromNat by.nat
%transform "Shift.inject" Shift.inject = Shift.injectViaNat
public export
shift : Shift from to -> Var from -> Var to
shift SZ i = i

View File

@ -299,3 +299,77 @@ mutual
pushSubstsWith th (comp th ps ph) e
pushSubstsWith th ph (DCloE e ps) =
pushSubstsWith (ps . th) ph e
parameters {0 d', n' : Nat}
mutual
namespace Term
export
inject : Term q d n -> Term q (d + d') (n + n')
inject (TYPE l) = TYPE l
inject (Pi qty arg res) = Pi qty (inject arg) (inject res)
inject (Lam body) = Lam (inject body)
inject (Sig fst snd) = Sig (inject fst) (inject snd)
inject (Pair fst snd) = Pair (inject fst) (inject snd)
inject (Enum cases) = Enum cases
inject (Tag tag) = Tag tag
inject (Eq ty l r) = Eq (inject ty) (inject l) (inject r)
inject (DLam body) = DLam (inject body)
inject (E e) = E (inject e)
inject (CloT tm th) = CloT (inject tm) (inject th)
inject (DCloT tm th) = DCloT (inject tm) (inject th)
namespace Elim
export
inject : Elim q d n -> Elim q (d + d') (n + n')
inject (F x) = F x
inject (B i) = B $ inject i
inject (fun :@ arg) = (inject fun) :@ (inject arg)
inject (CasePair qty pair ret body) =
CasePair qty (inject pair) (inject ret) (inject body)
inject (CaseEnum qty tag ret arms) =
CaseEnum qty (inject tag) (inject ret)
(assert_total $ map inject arms)
inject (fun :% arg) = (inject fun) :% (inject arg)
inject (tm :# ty) = (inject tm) :# (inject ty)
inject (CloE el th) = CloE (inject el) (inject th)
inject (DCloE el th) = DCloE (inject el) (inject th)
namespace ScopeTerm
export
inject : ScopeTermN s q d n -> ScopeTermN s q (d + d') (n + n')
inject (S names (N body)) = S names (N (inject body))
inject (S names (Y body)) {s, n} =
S names $ Y $ rewrite plusAssociative s n n' in inject body
namespace DScopeTerm
export
inject : DScopeTermN s q d n -> DScopeTermN s q (d + d') (n + n')
inject (S names (N body)) = S names (N (inject body))
inject (S names (Y body)) {s, d} =
S names $ Y $ rewrite plusAssociative s d d' in inject body
namespace TSubst
export
inject : TSubst q d from to -> TSubst q (d + d') (from + n') (to + n')
inject (Shift by) = Shift $ inject by
inject (t ::: th) = inject t ::: inject th
namespace DSubst
export
inject : DSubst from to -> DSubst (from + d') (to + d')
inject (Shift by) = Shift $ inject by
inject (p ::: th) = inject p ::: inject th
-- [fixme]
-- Error: Linearity checking failed on metavar Quox.Syntax.Term.Subst.{b:7362}
-- (Int not a function type)
{-
%transform "Term.inject" Term.inject t = believe_me t
%transform "Elim.inject" Elim.inject e = believe_me e
%transform "ScopeTerm.inject" ScopeTerm.inject t = believe_me t
%transform "DScopeTerm.inject" DScopeTerm.inject t = believe_me t
%transform "TSubst.inject" TSubst.inject th = believe_me th
%transform "DSubst.inject" DSubst.inject th = believe_me th
-}

View File

@ -111,6 +111,18 @@ toFromNat 0 (LTESucc x) = Refl
toFromNat (S k) (LTESucc x) = cong S $ toFromNat k x
export
inject : Var m -> Var (m + n)
inject VZ = VZ
inject (VS i) = VS $ inject i
private
injectViaNat : Var m -> Var (m + n)
injectViaNat i = fromNatWith i.nat (toNatLT i `transitive` lteAddRight m)
%transform "Var.inject" Var.inject = injectViaNat
-- not using %transform like other things because weakSpec requires the proof
-- to be relevant. but since only `LTESucc` is ever possible that seems
-- to be an instance of <https://github.com/idris-lang/Idris2/issues/1259>?

View File

@ -54,10 +54,10 @@ tests = "lexer" :- [
"identifiers & keywords" :- [
lexes "abc" [I "abc"],
lexes "abc def" [I "abc", I "def"],
lexes "abc def" [I "abc", R "def"],
lexes "abc_def" [I "abc_def"],
lexes "abc-def" [I "abc-def"],
lexes "abc{-comment-}def" [I "abc", I "def"],
lexes "abc{-comment-}def" [I "abc", R "def"],
lexes "λ" [R "λ"],
lexes "fun" [R "λ"],
lexes "δ" [R "δ"],

View File

@ -237,5 +237,25 @@ tests = "parser" :- [
Case Any (V "t") (Nothing, V "A") (CaseEnum []),
parsesAs term "case# t return A of {}" $
Case Any (V "t") (Nothing, V "A") (CaseEnum [])
],
"definitions" :- [
parsesAs definition "defω x : (_: {a}) × {b} ≔ ('a, 'b);" $
MkPDef Any "x" (Sig Nothing (Enum ["a"]) (Enum ["b"]))
(Pair (Tag "a") (Tag "b")),
parsesAs definition "def# x : (_: {a}) ** {b} ≔ ('a, 'b);" $
MkPDef Any "x" (Sig Nothing (Enum ["a"]) (Enum ["b"]))
(Pair (Tag "a") (Tag "b")),
parsesAs definition "def ω·x : (_: {a}) × {b} ≔ ('a, 'b);" $
MkPDef Any "x" (Sig Nothing (Enum ["a"]) (Enum ["b"]))
(Pair (Tag "a") (Tag "b")),
parsesAs definition "def x : (_: {a}) × {b} ≔ ('a, 'b);" $
MkPDef Any "x" (Sig Nothing (Enum ["a"]) (Enum ["b"]))
(Pair (Tag "a") (Tag "b")),
parsesAs definition "def0 A : ★₀ ≔ {a, b, c};" $
MkPDef Zero "A" (TYPE 0) (Enum ["a", "b", "c"]),
parsesAs input "def0 A : ★₀ ≔ {}; def0 B : ★₁ ≔ A;" $
[MkPDef Zero "A" (TYPE 0) (Enum []),
MkPDef Zero "B" (TYPE 1) (V "A")]
]
]