quox/lib/Quox/Parser/Parser.idr

331 lines
8.3 KiB
Idris
Raw Normal View History

module Quox.Parser.Parser
import public Quox.Parser.Lexer
import public Quox.Parser.Syntax
import Data.Fin
import Data.Vect
import public Text.Parser
%default total
public export
0 Grammar : Bool -> Type -> Type
Grammar = Core.Grammar () Token
%hide Core.Grammar
export
res : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
res str = terminal "expecting \"\{str}\"" $ guard . (== Reserved str)
export
resC : (str : String) -> (0 _ : IsReserved str) => Grammar True ()
resC str = do res str; commit
parameters {default True commit : Bool}
private
maybeCommit : Grammar False ()
maybeCommit = when commit Core.commit
export
betweenR : {c : Bool} -> (op, cl : String) ->
(0 _ : IsReserved op) => (0 _ : IsReserved cl) =>
Grammar c a -> Grammar True a
betweenR o c p = res o *> maybeCommit *> p <* res c <* maybeCommit
export
parens, bracks, braces : {c : Bool} -> Grammar c a -> Grammar True a
parens = betweenR "(" ")"
bracks = betweenR "[" "]"
braces = betweenR "{" "}"
export
commaSep : {c : Bool} -> Grammar c a -> Grammar False (List a)
commaSep = sepEndBy (res ",")
-- don't commit because of the possible terminating ","
export
semiSep : {c : Bool} -> Grammar c a -> Grammar False (List a)
semiSep = sepEndBy (res ";")
export
commaSep1 : {c : Bool} -> Grammar c a -> Grammar c (List1 a)
commaSep1 = sepEndBy1 (res ",")
export
darr : Grammar True ()
darr = resC ""
export
name : Grammar True PName
name = terminal "expecting name" $
\case Name i => Just i; _ => Nothing
export
mods : Grammar True Mods
mods = name <&> \n => n.mods :< n.base
export
baseName : Grammar True String
baseName = terminal "expecting unqualified name" $
\case Name i => guard (null i.mods) $> i.base
_ => Nothing
export
nat : Grammar True Nat
nat = terminal "expecting natural number" $
\case Nat n => pure n; _ => Nothing
export
string : Grammar True String
string = terminal "expecting string literal" $
\case Str s => pure s; _ => Nothing
export
tag : Grammar True String
tag = terminal "expecting tag constructor" $
\case Tag t => pure t; _ => Nothing
export
bareTag : Grammar True String
bareTag = string <|> [|toDotsP name|]
export
universe : Grammar True PUniverse
universe = terminal "expecting type universe" $
\case TYPE u => Just u; _ => Nothing
export
bname : Grammar True BName
bname = Nothing <$ res "_"
<|> [|Just baseName|]
export
dimConst : Grammar True DimConst
dimConst = terminal "expecting 0 or 1" $
\case Nat 0 => Just Zero; Nat 1 => Just One; _ => Nothing
export
qty : Grammar True PQty
qty = terminal "expecting quantity" $
\case Nat 0 => Just Zero; Nat 1 => Just One; Reserved "ω" => Just Any
_ => 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
qtys : Grammar False (List PQty)
qtys = option [] [|toList $ some qty <* res "."|]
export
dim : Grammar True PDim
dim = [|V baseName|] <|> [|K dimConst|]
private
0 PBinderHead : Nat -> Type
PBinderHead n = (Vect n PQty, BName, PTerm)
private
toVect : List a -> (n ** Vect n a)
toVect [] = (_ ** [])
toVect (x :: xs) = (_ ** x :: snd (toVect xs))
private
0 MakeBinder : Nat -> Type
MakeBinder n = (String, PBinderHead n -> PTerm -> PTerm)
private
makePi : MakeBinder 1
makePi = ("", \([pi], x, s) => Pi pi x s)
private
makeSig : MakeBinder 0
makeSig = ("×", \([], x, s) => Sig x s)
private
makeBinder : (m ** PBinderHead m) -> (n ** MakeBinder n) -> PTerm ->
Grammar False PTerm
makeBinder (m ** h) (n ** (str, f)) t =
case decEq m n of
Yes Refl => pure $ f h t
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)
binderInfix = symbols [((1 ** makePi), ""), ((0 ** makeSig), "×")]
private
lamIntro : Grammar True (BName -> PTerm -> PTerm)
lamIntro = symbolsC [(Lam, "λ"), (DLam, "δ")]
private covering
caseIntro : Grammar True PQty
caseIntro = symbols [(One, "case1"), (Any, "caseω")]
<|> resC "case" *> qty <* resC "."
private
optNameBinder : Grammar False BName
optNameBinder = [|join $ optional $ bname <* darr|]
mutual
export covering
term : Grammar True PTerm
term = lamTerm <|> caseTerm <|> bindTerm <|> annTerm
private covering
lamTerm : Grammar True PTerm
lamTerm = flip . foldr <$> lamIntro <*> some bname <* darr <*> term
private covering
caseTerm : Grammar True PTerm
caseTerm =
[|Case caseIntro term
(resC "return" *> optBinderTerm)
(resC "of" *> caseBody)|]
private covering
caseBody : Grammar True PCaseBody
caseBody = braces $
[|CasePair (pairPat <* darr) (term <* optSemi)|]
<|> [|CaseEnum $ semiSep [|MkPair tag (darr *> term)|]|]
where
optSemi = optional $ res ";"
pairPat = parens [|MkPair bname (resC "," *> bname)|]
private covering
bindTerm : Grammar True PTerm
bindTerm = join [|makeBinder binderHead binderInfix term|]
private covering
annTerm : Grammar True PTerm
annTerm = [|annotate infixEqTerm (optional $ resC "" *> term)|]
where annotate s a = maybe s (s :#) a
private covering
infixEqTerm : Grammar True PTerm
infixEqTerm = do
l <- infixTimesTerm
rty <- optional [|MkPair (resC "" *> term) (resC ":" *> infixTimesTerm)|]
pure $ maybe l (\rty => Eq (Nothing, snd rty) l (fst rty)) rty
private covering
infixTimesTerm : Grammar True PTerm
infixTimesTerm = foldr1 (Sig Nothing) <$> sepBy1 (resC "×") appTerm
private covering
appTerm : Grammar True PTerm
appTerm = resC "" *> [|TYPE nat|]
<|> resC "Eq" *> [|Eq (bracks optBinderTerm) aTerm aTerm|]
<|> [|apply aTerm (many appArg)|]
where
data PArg = TermArg PTerm | DimArg PDim
appArg : Grammar True PArg
appArg = [|DimArg $ resC "@" *> dim|] <|> [|TermArg aTerm|]
apply : PTerm -> List PArg -> PTerm
apply = foldl apply1 where
apply1 : PTerm -> PArg -> PTerm
apply1 f (TermArg x) = f :@ x
apply1 f (DimArg p) = f :% p
private covering
aTerm : Grammar True PTerm
aTerm = [|Enum $ braces $ commaSep bareTag|]
<|> [|TYPE universe|]
<|> [|V name|]
<|> [|Tag tag|]
<|> foldr1 Pair <$> parens (commaSep1 term)
private covering
binderHead : Grammar True (n ** PBinderHead n)
binderHead = parens {commit = False} $ do
qs <- [|toVect qtys|]
name <- bname
ty <- resC ":" *> term
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 = symbols [(Zero, "def0"), (Any, "defω")]
<|> resC "def" *> option Any (qty <* resC ".")
export covering
definition : Grammar True PDefinition
definition =
[|MkPDef defIntro name (optional (resC ":" *> term)) (resC "=" *> term)|]
export
load : Grammar True String
load = resC "load" *> string
mutual
export covering
namespace_ : Grammar True PNamespace
namespace_ =
[|MkPNamespace (resC "namespace" *> mods) (braces $ many decl)|]
export covering
decl : Grammar True PDecl
decl = [|PDef definition|]
<|> [|PNs namespace_|]
export covering
topLevel : Grammar True PTopLevel
topLevel = [|PLoad load|]
<|> [|PD decl|]
export covering
input : Grammar False (List PTopLevel)
input = skipSemis *> many (topLevel <* skipSemis)
where skipSemis = ignore $ many $ resC ";"
public export
data Error =
LexError Lexer.Error
| ParseError (List1 (ParsingError Token))
%hide Lexer.Error
export
lexParseWith : {c : Bool} -> Grammar c a -> String -> Either Error a
lexParseWith grm input = do
toks <- mapFst LexError $ lex input
bimap ParseError fst $ parse (grm <* eof) toks
export covering %inline
lexParseInput : String -> Either Error (List PTopLevel)
lexParseInput = lexParseWith input
export covering %inline
lexParseTerm : String -> Either Error PTerm
lexParseTerm = lexParseWith term