add file locations to Parser.Syntax
they're immediately thrown away currently. but one step at a time
This commit is contained in:
parent
97f51b4436
commit
7e079a9668
7 changed files with 822 additions and 667 deletions
|
@ -26,12 +26,27 @@ data Error =
|
|||
%hide Lexer.Error
|
||||
%runElab derive "Error" [Show]
|
||||
|
||||
|
||||
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
|
||||
withLoc : {c : Bool} -> FileName -> (Grammar c (Loc -> a)) -> Grammar c a
|
||||
withLoc fname act = bounds act <&> \res =>
|
||||
if res.isIrrelevant then res.val Nothing
|
||||
else res.val $ makeLoc fname res.bounds
|
||||
|
||||
export
|
||||
defLoc : FileName -> (Loc -> a) -> Grammar False a
|
||||
defLoc fname f = position <&> f . makeLoc fname
|
||||
|
||||
export
|
||||
unused : FileName -> Grammar False PatVar
|
||||
unused fname = defLoc fname Unused
|
||||
|
||||
|
||||
||| reserved token, like punctuation or keywords etc
|
||||
export
|
||||
|
@ -114,26 +129,33 @@ dimConst = terminalMatchN "dimension constant"
|
|||
|
||||
||| quantity (0, 1, or ω)
|
||||
export
|
||||
qty : Grammar True Qty
|
||||
qty = terminalMatchN "quantity"
|
||||
qtyVal : Grammar True Qty
|
||||
qtyVal = terminalMatchN "quantity"
|
||||
[(`(Nat 0), `(Zero)), (`(Nat 1), `(One)), (`(Reserved "ω"), `(Any))]
|
||||
|
||||
|
||||
||| quantity (0, 1, or ω)
|
||||
export
|
||||
qty : FileName -> Grammar True PQty
|
||||
qty fname = withLoc fname [|PQ qtyVal|]
|
||||
|
||||
|
||||
||| pattern var (unqualified name or _)
|
||||
export
|
||||
patVar : Grammar True BName
|
||||
patVar = [|Just baseName|] <|> Nothing <$ res "_"
|
||||
patVar : FileName -> Grammar True PatVar
|
||||
patVar fname = withLoc fname $
|
||||
[|PV baseName|] <|> Unused <$ res "_"
|
||||
|
||||
|
||||
||| dimension (without `@` prefix)
|
||||
export
|
||||
dim : Grammar True PDim
|
||||
dim = [|K dimConst|] <|> [|V baseName|]
|
||||
dim : FileName -> Grammar True PDim
|
||||
dim fname = withLoc fname $ [|K dimConst|] <|> [|V baseName|]
|
||||
|
||||
||| dimension argument (with @)
|
||||
export
|
||||
dimArg : Grammar True PDim
|
||||
dimArg = resC "@" *> mustWork dim
|
||||
dimArg : FileName -> Grammar True PDim
|
||||
dimArg fname = do resC "@"; mustWork $ dim fname
|
||||
|
||||
|
||||
delim : (o, c : String) -> (0 _ : IsReserved o) => (0 _ : IsReserved c) =>
|
||||
|
@ -168,313 +190,369 @@ enumType = delimSep "{" "}" "," bareTag
|
|||
|
||||
||| e.g. `case` or `case 1.`
|
||||
export
|
||||
caseIntro : Grammar True Qty
|
||||
caseIntro =
|
||||
Zero <$ res "case0"
|
||||
<|> One <$ res "case1"
|
||||
<|> Any <$ res "caseω"
|
||||
<|> delim "case" "." qty
|
||||
caseIntro : FileName -> Grammar True PQty
|
||||
caseIntro fname =
|
||||
withLoc fname (PQ Zero <$ res "case0")
|
||||
<|> withLoc fname (PQ One <$ res "case1")
|
||||
<|> withLoc fname (PQ Any <$ res "caseω")
|
||||
<|> delim "case" "." (qty fname)
|
||||
|
||||
export
|
||||
qtyPatVar : Grammar True (Qty, BName)
|
||||
qtyPatVar = [|(,) qty (needRes "." *> patVar)|]
|
||||
qtyPatVar : FileName -> Grammar True (PQty, PatVar)
|
||||
qtyPatVar fname = [|(,) (qty fname) (needRes "." *> patVar fname)|]
|
||||
|
||||
|
||||
export
|
||||
ptag : FileName -> Grammar True PTagVal
|
||||
ptag fname = withLoc fname $ [|PT tag|]
|
||||
|
||||
public export
|
||||
data PCasePat =
|
||||
PPair BName BName
|
||||
| PTag TagVal
|
||||
| PZero
|
||||
| PSucc BName Qty BName
|
||||
| PBox BName
|
||||
PPair PatVar PatVar Loc
|
||||
| PTag PTagVal Loc
|
||||
| PZero Loc
|
||||
| PSucc PatVar PQty PatVar Loc
|
||||
| PBox PatVar Loc
|
||||
%runElab derive "PCasePat" [Eq, Ord, Show]
|
||||
|
||||
export
|
||||
Located PCasePat where
|
||||
(PPair _ _ loc).loc = loc
|
||||
(PTag _ loc).loc = loc
|
||||
(PZero loc).loc = loc
|
||||
(PSucc _ _ _ loc).loc = loc
|
||||
(PBox _ loc).loc = loc
|
||||
|
||||
||| either `zero` or `0`
|
||||
export
|
||||
zeroPat : Grammar True ()
|
||||
zeroPat = resC "zero" <|> terminal "expected '0'" (guard . (== Nat 0))
|
||||
|
||||
export
|
||||
casePat : Grammar True PCasePat
|
||||
casePat =
|
||||
delim "(" ")" [|PPair patVar (needRes "," *> patVar)|]
|
||||
<|> [|PTag tag|]
|
||||
casePat : FileName -> Grammar True PCasePat
|
||||
casePat fname = withLoc fname $
|
||||
delim "(" ")" [|PPair (patVar fname) (needRes "," *> patVar fname)|]
|
||||
<|> [|PTag (ptag fname)|]
|
||||
<|> PZero <$ zeroPat
|
||||
<|> do p <- resC "succ" *> patVar
|
||||
ih <- option (Zero, Nothing) $ resC "," *> qtyPatVar
|
||||
<|> do p <- resC "succ" *> patVar fname
|
||||
ih <- resC "," *> qtyPatVar fname
|
||||
<|> [|(,) (defLoc fname $ PQ Zero) (unused fname)|]
|
||||
pure $ PSucc p (fst ih) (snd ih)
|
||||
<|> delim "[" "]" [|PBox patVar|]
|
||||
<|> delim "[" "]" [|PBox (patVar fname)|]
|
||||
<|> fatalError "invalid pattern"
|
||||
|
||||
export covering
|
||||
term : FileName -> Grammar True PTerm
|
||||
-- defined after all the subterm parsers
|
||||
|
||||
export covering
|
||||
term : Grammar True PTerm
|
||||
|
||||
export covering
|
||||
typeLine : Grammar True (BName, PTerm)
|
||||
typeLine = do
|
||||
typeLine : FileName -> Grammar True (PatVar, PTerm)
|
||||
typeLine fname = do
|
||||
resC "["
|
||||
mustWork $ do
|
||||
i <- option Nothing $ patVar <* resC "⇒"
|
||||
t <- term <* needRes "]"
|
||||
i <- patVar fname <* resC "⇒" <|> unused fname
|
||||
t <- term fname <* needRes "]"
|
||||
pure (i, t)
|
||||
|
||||
||| box term `[t]` or type `[π.A]`
|
||||
export covering
|
||||
boxTerm : Grammar True PTerm
|
||||
boxTerm = do
|
||||
boxTerm : FileName -> Grammar True PTerm
|
||||
boxTerm fname = withLoc fname $ do
|
||||
res "["; commit
|
||||
q <- optional $ qty <* res "."
|
||||
t <- mustWork $ term <* needRes "]"
|
||||
q <- optional $ qty fname <* res "."
|
||||
t <- mustWork $ term fname <* needRes "]"
|
||||
pure $ maybe (Box t) (\q => BOX q t) q
|
||||
|
||||
||| tuple term like `(a, b)`, or parenthesised single term.
|
||||
||| allows terminating comma. more than two elements are nested on the right:
|
||||
||| `(a, b, c, d) = (a, (b, (c, d)))`.
|
||||
export covering
|
||||
tupleTerm : Grammar True PTerm
|
||||
tupleTerm = foldr1 Pair <$> delimSep1 "(" ")" "," term
|
||||
tupleTerm : FileName -> Grammar True PTerm
|
||||
tupleTerm fname = withLoc fname $ do
|
||||
terms <- delimSep1 "(" ")" "," $ term fname
|
||||
pure $ \loc => foldr1 (\s, t => Pair s t loc) terms
|
||||
|
||||
||| argument/atomic term: single-token terms, or those with delimiters e.g.
|
||||
||| `[t]`
|
||||
export covering
|
||||
termArg : Grammar True PTerm
|
||||
termArg = [|TYPE universe1|]
|
||||
<|> [|Enum enumType|]
|
||||
<|> [|Tag tag|]
|
||||
<|> boxTerm
|
||||
<|> Nat <$ res "ℕ"
|
||||
<|> Zero <$ res "zero"
|
||||
<|> [|fromNat nat|]
|
||||
<|> [|V qname|]
|
||||
<|> tupleTerm
|
||||
termArg : FileName -> Grammar True PTerm
|
||||
termArg fname = withLoc fname $
|
||||
[|TYPE universe1|]
|
||||
<|> [|Enum enumType|]
|
||||
<|> [|Tag tag|]
|
||||
<|> const <$> boxTerm fname
|
||||
<|> Nat <$ res "ℕ"
|
||||
<|> Zero <$ res "zero"
|
||||
<|> [|fromNat nat|]
|
||||
<|> [|V qname|]
|
||||
<|> const <$> tupleTerm fname
|
||||
|
||||
export covering
|
||||
coeTerm : Grammar True PTerm
|
||||
coeTerm = resC "coe" *> mustWork [|Coe typeLine dimArg dimArg termArg|]
|
||||
coeTerm : FileName -> Grammar True PTerm
|
||||
coeTerm fname = withLoc fname $ do
|
||||
resC "coe"
|
||||
mustWork [|Coe (typeLine fname) (dimArg fname) (dimArg fname)
|
||||
(termArg fname)|]
|
||||
|
||||
public export
|
||||
CompBranch : Type
|
||||
CompBranch = (DimConst, PatVar, PTerm)
|
||||
|
||||
export covering
|
||||
compBranch : Grammar True (DimConst, BName, PTerm)
|
||||
compBranch = [|(,,) dimConst patVar (needRes "⇒" *> term)|]
|
||||
compBranch : FileName -> Grammar True CompBranch
|
||||
compBranch fname = [|(,,) dimConst (patVar fname) (needRes "⇒" *> term fname)|]
|
||||
|
||||
private
|
||||
checkCompTermBody : (PatVar, PTerm) -> PDim -> PDim -> PTerm -> PDim ->
|
||||
CompBranch -> CompBranch -> Bounds ->
|
||||
Grammar False (Loc -> PTerm)
|
||||
checkCompTermBody a p q s r (e0, s0) (e1, s1) bounds =
|
||||
case (e0, e1) of
|
||||
(Zero, One) => pure $ Comp a p q s r s0 s1
|
||||
(One, Zero) => pure $ Comp a p q s r s1 s0
|
||||
(_, _) =>
|
||||
fatalLoc bounds "body of 'comp' needs one 0 case and one 1 case"
|
||||
|
||||
export covering
|
||||
compTerm : Grammar True PTerm
|
||||
compTerm = do
|
||||
compTerm : FileName -> Grammar True PTerm
|
||||
compTerm fname = withLoc fname $ do
|
||||
resC "comp"
|
||||
mustWork $ do
|
||||
a <- typeLine
|
||||
p <- dimArg; q <- dimArg
|
||||
s <- termArg; r <- dimArg
|
||||
needRes "{"
|
||||
s0 <- compBranch; needRes ";"
|
||||
s1 <- compBranch; optRes ";"
|
||||
needRes "}"
|
||||
case (fst s0, fst s1) of
|
||||
(Zero, One) => pure $ Comp a p q s r (snd s0) (snd s1)
|
||||
(One, Zero) => pure $ Comp a p q s r (snd s1) (snd s0)
|
||||
(_, _) =>
|
||||
fatalError "body of 'comp' needs one 0 case and one 1 case"
|
||||
a <- typeLine fname
|
||||
p <- dimArg fname; q <- dimArg fname
|
||||
s <- termArg fname; r <- dimArg fname
|
||||
bodyStart <- bounds $ needRes "{"
|
||||
s0 <- compBranch fname; needRes ";"
|
||||
s1 <- compBranch fname; optRes ";"
|
||||
bodyEnd <- bounds $ needRes "}"
|
||||
let body = bounds $ mergeBounds bodyStart bodyEnd
|
||||
checkCompTermBody a p q s r s0 s1 body
|
||||
|
||||
export covering
|
||||
splitUniverseTerm : Grammar True PTerm
|
||||
splitUniverseTerm = resC "★" *> mustWork [|TYPE nat|]
|
||||
splitUniverseTerm : FileName -> Grammar True PTerm
|
||||
splitUniverseTerm fname = withLoc fname $ resC "★" *> mustWork [|TYPE nat|]
|
||||
|
||||
export covering
|
||||
eqTerm : Grammar True PTerm
|
||||
eqTerm = resC "Eq" *> mustWork [|Eq typeLine termArg termArg|]
|
||||
eqTerm : FileName -> Grammar True PTerm
|
||||
eqTerm fname = withLoc fname $
|
||||
resC "Eq" *> mustWork [|Eq (typeLine fname) (termArg fname) (termArg fname)|]
|
||||
|
||||
export covering
|
||||
succTerm : Grammar True PTerm
|
||||
succTerm = resC "succ" *> mustWork [|Succ termArg|]
|
||||
succTerm : FileName -> Grammar True PTerm
|
||||
succTerm fname = withLoc fname $
|
||||
resC "succ" *> mustWork [|Succ (termArg fname)|]
|
||||
|
||||
||| a dimension argument with an `@` prefix, or
|
||||
||| a term argument with no prefix
|
||||
export covering
|
||||
anyArg : Grammar True (Either PDim PTerm)
|
||||
anyArg = dimArg <||> termArg
|
||||
anyArg : FileName -> Grammar True (Either PDim PTerm)
|
||||
anyArg fname = dimArg fname <||> termArg fname
|
||||
|
||||
export covering
|
||||
normalAppTerm : Grammar True PTerm
|
||||
normalAppTerm = foldl ap <$> termArg <*> many anyArg
|
||||
where ap : PTerm -> Either PDim PTerm -> PTerm
|
||||
ap f (Left p) = f :% p
|
||||
ap f (Right s) = f :@ s
|
||||
normalAppTerm : FileName -> Grammar True PTerm
|
||||
normalAppTerm fname = withLoc fname $ do
|
||||
head <- termArg fname
|
||||
args <- many $ anyArg fname
|
||||
pure $ \loc => foldl (ap loc) head args
|
||||
where ap : Loc -> PTerm -> Either PDim PTerm -> PTerm
|
||||
ap loc f (Left p) = DApp f p loc
|
||||
ap loc f (Right s) = App f s loc
|
||||
|
||||
||| application term `f x @y z`, or other terms that look like application
|
||||
||| like `succ` or `coe`.
|
||||
export covering
|
||||
appTerm : Grammar True PTerm
|
||||
appTerm = coeTerm
|
||||
<|> compTerm
|
||||
<|> splitUniverseTerm
|
||||
<|> eqTerm
|
||||
<|> succTerm
|
||||
<|> normalAppTerm
|
||||
appTerm : FileName -> Grammar True PTerm
|
||||
appTerm fname =
|
||||
coeTerm fname
|
||||
<|> compTerm fname
|
||||
<|> splitUniverseTerm fname
|
||||
<|> eqTerm fname
|
||||
<|> succTerm fname
|
||||
<|> normalAppTerm fname
|
||||
|
||||
export covering
|
||||
infixEqTerm : Grammar True PTerm
|
||||
infixEqTerm = do
|
||||
l <- appTerm; commit
|
||||
rest <- optional $ res "≡" *> [|(,) term (needRes ":" *> appTerm)|]
|
||||
pure $ maybe l (\rest => Eq (Nothing, snd rest) l (fst rest)) rest
|
||||
infixEqTerm : FileName -> Grammar True PTerm
|
||||
infixEqTerm fname = withLoc fname $ do
|
||||
l <- appTerm fname; commit
|
||||
rest <- optional $
|
||||
res "≡" *> [|(,) (term fname) (needRes ":" *> appTerm fname)|]
|
||||
let u = Unused $ onlyStart l.loc
|
||||
pure $ \loc => maybe l (\rest => Eq (u, snd rest) l (fst rest) loc) rest
|
||||
|
||||
export covering
|
||||
annTerm : Grammar True PTerm
|
||||
annTerm = do
|
||||
tm <- infixEqTerm; commit
|
||||
ty <- optional $ res "∷" *> term
|
||||
pure $ maybe tm (tm :#) ty
|
||||
annTerm : FileName -> Grammar True PTerm
|
||||
annTerm fname = withLoc fname $ do
|
||||
tm <- infixEqTerm fname; commit
|
||||
ty <- optional $ res "∷" *> term fname
|
||||
pure $ \loc => maybe tm (\ty => Ann tm ty loc) ty
|
||||
|
||||
export covering
|
||||
lamTerm : Grammar True PTerm
|
||||
lamTerm = do
|
||||
lamTerm : FileName -> Grammar True PTerm
|
||||
lamTerm fname = withLoc fname $ do
|
||||
k <- DLam <$ res "δ" <|> Lam <$ res "λ"
|
||||
mustWork $ do
|
||||
xs <- some patVar; needRes "⇒"
|
||||
body <- term; commit
|
||||
pure $ foldr k body xs
|
||||
xs <- some $ patVar fname; needRes "⇒"
|
||||
body <- term fname; commit
|
||||
pure $ \loc => foldr (\x, s => k x s loc) body xs
|
||||
|
||||
-- [todo] fix the backtracking in e.g. (F x y z × B)
|
||||
export covering
|
||||
properBinders : Grammar True (List1 BName, PTerm)
|
||||
properBinders = do
|
||||
properBinders : FileName -> Grammar True (List1 PatVar, PTerm)
|
||||
properBinders fname = do
|
||||
res "("
|
||||
xs <- some patVar; resC ":"
|
||||
t <- term; needRes ")"
|
||||
xs <- some $ patVar fname; resC ":"
|
||||
t <- term fname; needRes ")"
|
||||
pure (xs, t)
|
||||
|
||||
export covering
|
||||
piTerm : Grammar True PTerm
|
||||
piTerm = do
|
||||
q <- qty; resC "."
|
||||
dom <- piBinder; needRes "→"
|
||||
cod <- term; commit
|
||||
pure $ foldr (\x => Pi q x (snd dom)) cod (fst dom)
|
||||
piTerm : FileName -> Grammar True PTerm
|
||||
piTerm fname = withLoc fname $ do
|
||||
q <- qty fname; resC "."
|
||||
dom <- piBinder; needRes "→"
|
||||
cod <- term fname; commit
|
||||
pure $ \loc => foldr (\x, t => Pi q x (snd dom) t loc) cod (fst dom)
|
||||
where
|
||||
piBinder : Grammar True (List1 BName, PTerm)
|
||||
piBinder = properBinders <|> [|(singleton Nothing,) termArg|]
|
||||
piBinder : Grammar True (List1 PatVar, PTerm)
|
||||
piBinder = properBinders fname
|
||||
<|> [|(,) [|singleton $ unused fname|] (termArg fname)|]
|
||||
|
||||
export covering
|
||||
sigmaTerm : Grammar True PTerm
|
||||
sigmaTerm =
|
||||
(properBinders >>= continueDep)
|
||||
<|> (annTerm >>= continueNondep)
|
||||
sigmaTerm : FileName -> Grammar True PTerm
|
||||
sigmaTerm fname =
|
||||
(properBinders fname >>= continueDep)
|
||||
<|> (annTerm fname >>= continueNondep)
|
||||
where
|
||||
continueDep : (List1 BName, PTerm) -> Grammar True PTerm
|
||||
continueDep (names, dom) = do
|
||||
cod <- needRes "×" *> sigmaTerm
|
||||
pure $ foldr (\x => Sig x dom) cod names
|
||||
continueNondep : PTerm -> Grammar False PTerm
|
||||
continueNondep dom = do
|
||||
rest <- optional $ resC "×" *> sepBy1 (res "×") annTerm
|
||||
Core.pure $ foldr1 (Sig Nothing) $ dom ::: maybe [] forget rest
|
||||
continueDep : (List1 PatVar, PTerm) -> Grammar True PTerm
|
||||
continueDep (names, fst) = withLoc fname $ do
|
||||
snd <- needRes "×" *> sigmaTerm fname
|
||||
pure $ \loc => foldr (\x, snd => Sig x fst snd loc) snd names
|
||||
|
||||
cross : PTerm -> PTerm -> PTerm
|
||||
cross l r = let loc = extend' l.loc r.loc.bounds in
|
||||
Sig (Unused $ onlyStart l.loc) l r loc
|
||||
|
||||
continueNondep : PTerm -> Grammar False PTerm
|
||||
continueNondep fst = do
|
||||
rest <- optional $ resC "×" *> sepBy1 (res "×") (annTerm fname)
|
||||
pure $ foldr1 cross $ fst ::: maybe [] toList rest
|
||||
|
||||
public export
|
||||
PCaseArm : Type
|
||||
PCaseArm = (PCasePat, PTerm)
|
||||
|
||||
export covering
|
||||
caseArm : FileName -> Grammar True PCaseArm
|
||||
caseArm fname = [|(,) (casePat fname) (needRes "⇒" *> term fname)|]
|
||||
|
||||
export
|
||||
checkCaseArms : List PCaseArm -> Grammar False PCaseBody
|
||||
checkCaseArms [] =
|
||||
pure $ CaseEnum []
|
||||
checkCaseArms ((PPair x y, rhs) :: rest) = do
|
||||
let [] = rest | _ => fatalError "unexpected pattern after pair"
|
||||
pure $ CasePair (x, y) rhs
|
||||
checkCaseArms ((PTag tag, rhs1) :: rest) = do
|
||||
checkCaseArms : Loc -> List PCaseArm -> Grammar False PCaseBody
|
||||
checkCaseArms loc [] = pure $ CaseEnum [] loc
|
||||
checkCaseArms loc ((PPair x y _, rhs) :: rest) =
|
||||
if null rest then pure $ CasePair (x, y) rhs loc
|
||||
else fatalError "unexpected pattern after pair"
|
||||
checkCaseArms loc ((PTag tag _, rhs1) :: rest) = do
|
||||
let rest = for rest $ \case
|
||||
(PTag tag, rhs) => Just (tag, rhs)
|
||||
_ => Nothing
|
||||
(PTag tag _, rhs) => Just (tag, rhs)
|
||||
_ => Nothing
|
||||
maybe (fatalError "expected all patterns to be tags")
|
||||
(pure . CaseEnum . ((tag, rhs1) ::)) rest
|
||||
checkCaseArms ((PZero, rhs1) :: rest) = do
|
||||
let [(PSucc p q ih, rhs2)] = rest
|
||||
(\rest => pure $ CaseEnum ((tag, rhs1) :: rest) loc) rest
|
||||
checkCaseArms loc ((PZero _, rhs1) :: rest) = do
|
||||
let [(PSucc p q ih _, rhs2)] = rest
|
||||
| _ => fatalError "expected succ pattern after zero"
|
||||
pure $ CaseNat rhs1 (p, q, ih, rhs2)
|
||||
checkCaseArms ((PSucc p q ih, rhs1) :: rest) = do
|
||||
let [(PZero, rhs2)] = rest
|
||||
pure $ CaseNat rhs1 (p, q, ih, rhs2) loc
|
||||
checkCaseArms loc ((PSucc p q ih _, rhs1) :: rest) = do
|
||||
let [(PZero _, rhs2)] = rest
|
||||
| _ => fatalError "expected zero pattern after succ"
|
||||
pure $ CaseNat rhs2 (p, q, ih, rhs1)
|
||||
checkCaseArms ((PBox x, rhs) :: rest) = do
|
||||
let [] = rest | _ => fatalError "unexpected pattern after box"
|
||||
pure $ CaseBox x rhs
|
||||
pure $ CaseNat rhs2 (p, q, ih, rhs1) loc
|
||||
checkCaseArms loc ((PBox x _, rhs) :: rest) =
|
||||
if null rest then pure $ CaseBox x rhs loc
|
||||
else fatalError "unexpected pattern after box"
|
||||
|
||||
export covering
|
||||
caseArm : Grammar True PCaseArm
|
||||
caseArm = [|(,) casePat (needRes "⇒" *> term)|]
|
||||
caseBody : FileName -> Grammar True PCaseBody
|
||||
caseBody fname = do
|
||||
body <- bounds $ delimSep "{" "}" ";" $ caseArm fname
|
||||
let loc = makeLoc fname body.bounds
|
||||
checkCaseArms loc body.val
|
||||
|
||||
export covering
|
||||
caseBody : Grammar True PCaseBody
|
||||
caseBody = delimSep "{" "}" ";" caseArm >>= checkCaseArms
|
||||
caseReturn : FileName -> Grammar True (PatVar, PTerm)
|
||||
caseReturn fname = do
|
||||
x <- patVar fname <* resC "⇒" <|> unused fname
|
||||
ret <- term fname
|
||||
pure (x, ret)
|
||||
|
||||
export covering
|
||||
caseTerm : Grammar True PTerm
|
||||
caseTerm = do
|
||||
qty <- caseIntro; commit
|
||||
head <- mustWork term; needRes "return"
|
||||
ret <- mustWork [|(,) (option Nothing $ patVar <* resC "⇒") term|]
|
||||
needRes "of"
|
||||
body <- mustWork caseBody
|
||||
caseTerm : FileName -> Grammar True PTerm
|
||||
caseTerm fname = withLoc fname $ do
|
||||
qty <- caseIntro fname; commit
|
||||
head <- mustWork $ term fname; needRes "return"
|
||||
ret <- mustWork $ caseReturn fname; needRes "of"
|
||||
body <- mustWork $ caseBody fname
|
||||
pure $ Case qty head ret body
|
||||
|
||||
-- export covering
|
||||
-- term : FileName -> Grammar True PTerm
|
||||
term fname = lamTerm fname
|
||||
<|> caseTerm fname
|
||||
<|> piTerm fname
|
||||
<|> sigmaTerm fname
|
||||
|
||||
term = lamTerm
|
||||
<|> caseTerm
|
||||
<|> piTerm
|
||||
<|> sigmaTerm
|
||||
|
||||
export covering
|
||||
decl : FileName -> Grammar True PDecl
|
||||
|
||||
||| `def` alone means `defω`
|
||||
export
|
||||
defIntro : Grammar True Qty
|
||||
defIntro = Zero <$ resC "def0"
|
||||
<|> Any <$ resC "defω"
|
||||
<|> resC "def" *> option Any (qty <* needRes ".")
|
||||
|
||||
defIntro : FileName -> Grammar True PQty
|
||||
defIntro fname =
|
||||
withLoc fname (PQ Zero <$ resC "def0")
|
||||
<|> withLoc fname (PQ Any <$ resC "defω")
|
||||
<|> do pos <- bounds $ resC "def"
|
||||
let any = PQ Any $ makeLoc fname pos.bounds
|
||||
option any $ qty fname <* needRes "."
|
||||
|
||||
export covering
|
||||
decl : Grammar True PDecl
|
||||
|
||||
export covering
|
||||
definition : Grammar True PDefinition
|
||||
definition = do
|
||||
qty <- defIntro
|
||||
definition : FileName -> Grammar True PDefinition
|
||||
definition fname = withLoc fname $ do
|
||||
qty <- defIntro fname
|
||||
name <- baseName
|
||||
type <- optional $ resC ":" *> mustWork term
|
||||
term <- needRes "=" *> mustWork term
|
||||
type <- optional $ resC ":" *> mustWork (term fname)
|
||||
term <- needRes "=" *> mustWork (term fname)
|
||||
optRes ";"
|
||||
pure $ MkPDef {qty, name, type, term}
|
||||
pure $ MkPDef qty name type term
|
||||
|
||||
export covering
|
||||
namespace_ : Grammar True PNamespace
|
||||
namespace_ = do
|
||||
namespace_ : FileName -> Grammar True PNamespace
|
||||
namespace_ fname = withLoc fname $ do
|
||||
ns <- resC "namespace" *> qname; needRes "{"
|
||||
decls <- nsInner; optRes ";"
|
||||
pure $ MkPNamespace (ns.mods :< ns.base) decls
|
||||
where
|
||||
nsInner : Grammar True (List PDecl)
|
||||
nsInner = [] <$ resC "}"
|
||||
<|> [|(decl <* commit) :: nsInner|]
|
||||
<|> [|(decl fname <* commit) :: nsInner|]
|
||||
|
||||
decl = [|PDef definition|] <|> [|PNs namespace_|]
|
||||
decl fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
||||
|
||||
export covering
|
||||
load : Grammar True String
|
||||
load = resC "load" *> mustWork strLit <* optRes ";"
|
||||
load : FileName -> Grammar True PTopLevel
|
||||
load fname = withLoc fname $
|
||||
resC "load" *> mustWork [|PLoad strLit|] <* optRes ";"
|
||||
|
||||
export covering
|
||||
topLevel : Grammar True PTopLevel
|
||||
topLevel = [|PLoad load|] <|> [|PD decl|]
|
||||
topLevel : FileName -> Grammar True PTopLevel
|
||||
topLevel fname = load fname <|> [|PD $ decl fname|]
|
||||
|
||||
export covering
|
||||
input : Grammar False (List PTopLevel)
|
||||
input = [] <$ eof
|
||||
<|> [|(topLevel <* commit) :: input|]
|
||||
input : FileName -> Grammar False (List PTopLevel)
|
||||
input fname = [] <$ eof
|
||||
<|> [|(topLevel fname <* commit) :: input fname|]
|
||||
|
||||
export covering
|
||||
lexParseTerm : String -> Either Error PTerm
|
||||
lexParseTerm = lexParseWith term
|
||||
lexParseTerm : FileName -> String -> Either Error PTerm
|
||||
lexParseTerm = lexParseWith . term
|
||||
|
||||
export covering
|
||||
lexParseInput : String -> Either Error (List PTopLevel)
|
||||
lexParseInput = lexParseWith input
|
||||
lexParseInput : FileName -> String -> Either Error (List PTopLevel)
|
||||
lexParseInput = lexParseWith . input
|
||||
|
|
Loading…
Add table
Add a link
Reference in a new issue