parser tweaks
qtys and dims don't allow useless parens any more. everything else should be the same
This commit is contained in:
parent
51468f54fc
commit
ebf6aefb1d
2 changed files with 50 additions and 52 deletions
|
@ -18,8 +18,7 @@ Grammar = Core.Grammar () Token
|
||||||
|
|
||||||
export
|
export
|
||||||
res : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
|
res : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
|
||||||
res str = terminal "expecting \"\{str}\"" $
|
res str = terminal "expecting \"\{str}\"" $ guard . (== Reserved str)
|
||||||
\x => guard $ x == Reserved str
|
|
||||||
|
|
||||||
export
|
export
|
||||||
resC : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
|
resC : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
|
||||||
|
@ -107,30 +106,41 @@ bname = Nothing <$ res "_"
|
||||||
<|> [|Just baseName|]
|
<|> [|Just baseName|]
|
||||||
|
|
||||||
export
|
export
|
||||||
zeroOne : (zero, one : a) -> Grammar True a
|
dimConst : Grammar True DimConst
|
||||||
zeroOne zero one = terminal "expecting zero or one" $
|
dimConst = terminal "expecting 0 or 1" $
|
||||||
\case Nat 0 => Just zero; Nat 1 => Just one; _ => Nothing
|
\case Nat 0 => Just Zero; Nat 1 => Just One; _ => Nothing
|
||||||
|
|
||||||
|
export
|
||||||
export covering
|
|
||||||
qty : Grammar True PQty
|
qty : Grammar True PQty
|
||||||
qty = zeroOne Zero One
|
qty = terminal "expecting quantity" $
|
||||||
<|> Any <$ res "ω"
|
\case Nat 0 => Just Zero; Nat 1 => Just One; Reserved "ω" => Just Any
|
||||||
<|> parens qty
|
_ => Nothing
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
AllReserved : List (a, String) -> Type
|
||||||
|
AllReserved = foldr (\x, ps => (IsReserved $ snd x, ps)) ()
|
||||||
|
|
||||||
|
export
|
||||||
|
symbols : (lst : List (a, String)) -> (0 _ : AllReserved lst) =>
|
||||||
|
Grammar True a
|
||||||
|
symbols [] = fail "no symbols found"
|
||||||
|
symbols ((x, str) :: rest) = x <$ res str <|> symbols rest
|
||||||
|
|
||||||
|
export
|
||||||
|
symbolsC : (lst : List (a, String)) -> (0 _ : AllReserved lst) =>
|
||||||
|
Grammar True a
|
||||||
|
symbolsC lst = symbols lst <* commit
|
||||||
|
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
qtys : Grammar False (List PQty)
|
qtys : Grammar False (List PQty)
|
||||||
qtys = option [] [|toList $ some qty <* res "."|]
|
qtys = option [] [|toList $ some qty <* res "."|]
|
||||||
|
|
||||||
export
|
|
||||||
dimConst : Grammar True DimConst
|
|
||||||
dimConst = zeroOne Zero One
|
|
||||||
|
|
||||||
export covering
|
export
|
||||||
dim : Grammar True PDim
|
dim : Grammar True PDim
|
||||||
dim = [|V baseName|]
|
dim = [|V baseName|] <|> [|K dimConst|]
|
||||||
<|> [|K dimConst|]
|
|
||||||
<|> parens dim
|
|
||||||
|
|
||||||
private
|
private
|
||||||
0 PBinderHead : Nat -> Type
|
0 PBinderHead : Nat -> Type
|
||||||
|
@ -154,9 +164,9 @@ makeSig : MakeBinder 0
|
||||||
makeSig = ("×", \([], x, s) => Sig x s)
|
makeSig = ("×", \([], x, s) => Sig x s)
|
||||||
|
|
||||||
private
|
private
|
||||||
makeBinder : {m, n : Nat} -> MakeBinder m -> PBinderHead n -> PTerm ->
|
makeBinder : (m ** PBinderHead m) -> (n ** MakeBinder n) -> PTerm ->
|
||||||
Grammar False PTerm
|
Grammar False PTerm
|
||||||
makeBinder (str, f) h t =
|
makeBinder (m ** h) (n ** (str, f)) t =
|
||||||
case decEq m n of
|
case decEq m n of
|
||||||
Yes Refl => pure $ f h t
|
Yes Refl => pure $ f h t
|
||||||
No _ =>
|
No _ =>
|
||||||
|
@ -165,18 +175,15 @@ makeBinder (str, f) h t =
|
||||||
|
|
||||||
private
|
private
|
||||||
binderInfix : Grammar True (n ** MakeBinder n)
|
binderInfix : Grammar True (n ** MakeBinder n)
|
||||||
binderInfix = res "→" $> (1 ** makePi)
|
binderInfix = symbols [((1 ** makePi), "→"), ((0 ** makeSig), "×")]
|
||||||
<|> res "×" $> (0 ** makeSig)
|
|
||||||
|
|
||||||
private
|
private
|
||||||
lamIntro : Grammar True (BName -> PTerm -> PTerm)
|
lamIntro : Grammar True (BName -> PTerm -> PTerm)
|
||||||
lamIntro = Lam <$ resC "λ"
|
lamIntro = symbolsC [(Lam, "λ"), (DLam, "δ")]
|
||||||
<|> DLam <$ resC "δ"
|
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
caseIntro : Grammar True PQty
|
caseIntro : Grammar True PQty
|
||||||
caseIntro = resC "case1" $> One
|
caseIntro = symbols [(One, "case1"), (Any, "caseω")]
|
||||||
<|> resC "caseω" $> Any
|
|
||||||
<|> resC "case" *> qty <* resC "."
|
<|> resC "case" *> qty <* resC "."
|
||||||
|
|
||||||
private
|
private
|
||||||
|
@ -186,13 +193,7 @@ optNameBinder = [|join $ optional $ bname <* darr|]
|
||||||
mutual
|
mutual
|
||||||
export covering
|
export covering
|
||||||
term : Grammar True PTerm
|
term : Grammar True PTerm
|
||||||
term = lamTerm
|
term = lamTerm <|> caseTerm <|> bindTerm <|> annTerm
|
||||||
<|> caseTerm
|
|
||||||
<|> bindTerm
|
|
||||||
<|> [|annotate infixEqTerm (optional $ resC "∷" *> term)|]
|
|
||||||
where
|
|
||||||
annotate : PTerm -> Maybe PTerm -> PTerm
|
|
||||||
annotate s a = maybe s (s :#) a
|
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
lamTerm : Grammar True PTerm
|
lamTerm : Grammar True PTerm
|
||||||
|
@ -203,23 +204,25 @@ mutual
|
||||||
caseTerm =
|
caseTerm =
|
||||||
[|Case caseIntro term
|
[|Case caseIntro term
|
||||||
(resC "return" *> optBinderTerm)
|
(resC "return" *> optBinderTerm)
|
||||||
(resC "of" *> braces caseBody)|]
|
(resC "of" *> caseBody)|]
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
caseBody : Grammar False PCaseBody
|
caseBody : Grammar True PCaseBody
|
||||||
caseBody = [|CasePair (pairPat <* darr) (term <* optSemi)|]
|
caseBody = braces $
|
||||||
|
[|CasePair (pairPat <* darr) (term <* optSemi)|]
|
||||||
<|> [|CaseEnum $ semiSep [|MkPair tag (darr *> term)|]|]
|
<|> [|CaseEnum $ semiSep [|MkPair tag (darr *> term)|]|]
|
||||||
where
|
where
|
||||||
optSemi = ignore $ optional $ res ";"
|
optSemi = optional $ res ";"
|
||||||
pairPat = parens [|MkPair bname (resC "," *> bname)|]
|
pairPat = parens [|MkPair bname (resC "," *> bname)|]
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
bindTerm : Grammar True PTerm
|
bindTerm : Grammar True PTerm
|
||||||
bindTerm = do
|
bindTerm = join [|makeBinder binderHead binderInfix term|]
|
||||||
bnd <- binderHead
|
|
||||||
inf <- binderInfix
|
private covering
|
||||||
body <- term
|
annTerm : Grammar True PTerm
|
||||||
makeBinder (snd inf) (snd bnd) body
|
annTerm = [|annotate infixEqTerm (optional $ resC "∷" *> term)|]
|
||||||
|
where annotate s a = maybe s (s :#) a
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
infixEqTerm : Grammar True PTerm
|
infixEqTerm : Grammar True PTerm
|
||||||
|
@ -241,8 +244,7 @@ mutual
|
||||||
data PArg = TermArg PTerm | DimArg PDim
|
data PArg = TermArg PTerm | DimArg PDim
|
||||||
|
|
||||||
appArg : Grammar True PArg
|
appArg : Grammar True PArg
|
||||||
appArg = [|DimArg $ resC "@" *> dim|]
|
appArg = [|DimArg $ resC "@" *> dim|] <|> [|TermArg aTerm|]
|
||||||
<|> [|TermArg aTerm|]
|
|
||||||
|
|
||||||
apply : PTerm -> List PArg -> PTerm
|
apply : PTerm -> List PArg -> PTerm
|
||||||
apply = foldl apply1 where
|
apply = foldl apply1 where
|
||||||
|
@ -263,8 +265,7 @@ mutual
|
||||||
binderHead = parens {commit = False} $ do
|
binderHead = parens {commit = False} $ do
|
||||||
qs <- [|toVect qtys|]
|
qs <- [|toVect qtys|]
|
||||||
name <- bname
|
name <- bname
|
||||||
resC ":"
|
ty <- resC ":" *> term
|
||||||
ty <- term
|
|
||||||
pure (qs.fst ** (qs.snd, name, ty))
|
pure (qs.fst ** (qs.snd, name, ty))
|
||||||
|
|
||||||
private covering
|
private covering
|
||||||
|
@ -274,8 +275,7 @@ mutual
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
defIntro : Grammar True PQty
|
defIntro : Grammar True PQty
|
||||||
defIntro = Zero <$ resC "def0"
|
defIntro = symbols [(Zero, "def0"), (Any, "defω")]
|
||||||
<|> Any <$ resC "defω"
|
|
||||||
<|> resC "def" *> option Any (qty <* resC ".")
|
<|> resC "def" *> option Any (qty <* resC ".")
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
|
|
|
@ -70,8 +70,7 @@ tests = "parser" :- [
|
||||||
"dimensions" :- [
|
"dimensions" :- [
|
||||||
parsesAs dim "0" (K Zero),
|
parsesAs dim "0" (K Zero),
|
||||||
parsesAs dim "1" (K One),
|
parsesAs dim "1" (K One),
|
||||||
parsesAs dim "ι" (V "ι"),
|
parsesAs dim "𝑖" (V "𝑖"),
|
||||||
parsesAs dim "(0)" (K Zero),
|
|
||||||
parseFails dim "M.x",
|
parseFails dim "M.x",
|
||||||
parseFails dim "_"
|
parseFails dim "_"
|
||||||
],
|
],
|
||||||
|
@ -81,7 +80,6 @@ tests = "parser" :- [
|
||||||
parsesAs qty "1" One,
|
parsesAs qty "1" One,
|
||||||
parsesAs qty "ω" Any,
|
parsesAs qty "ω" Any,
|
||||||
parsesAs qty "#" Any,
|
parsesAs qty "#" Any,
|
||||||
parsesAs qty "(#)" Any,
|
|
||||||
parseFails qty "anythingElse",
|
parseFails qty "anythingElse",
|
||||||
parseFails qty "_"
|
parseFails qty "_"
|
||||||
],
|
],
|
||||||
|
|
Loading…
Reference in a new issue