start of parser stuff
This commit is contained in:
parent
d43a2429e1
commit
79211cff84
8 changed files with 207 additions and 6 deletions
|
@ -5,7 +5,7 @@ authors = "rhiannon morris"
|
|||
sourceloc = "https://git.rhiannon.website/rhi/quox"
|
||||
license = "acsl"
|
||||
|
||||
depends = base, contrib, elab-util, sop
|
||||
depends = base, contrib, elab-util, sop, snocvect
|
||||
|
||||
sourcedir = "src"
|
||||
modules =
|
||||
|
@ -28,6 +28,7 @@ modules =
|
|||
Quox.Syntax.Var,
|
||||
Quox.Token,
|
||||
Quox.Lexer,
|
||||
Quox.Parser,
|
||||
Quox.Context,
|
||||
Quox.Equal,
|
||||
Quox.Name,
|
||||
|
|
|
@ -119,10 +119,10 @@ tokens = choice [
|
|||
|
||||
|
||||
export
|
||||
lex : MonadError Error m => String -> m (List (WithBounds Token))
|
||||
lex : String -> Either Error (List BToken)
|
||||
lex str =
|
||||
let (res, (reason, line, col, str)) = lex tokens str in
|
||||
case reason of
|
||||
EndInput => pure $ mapMaybe sequence res
|
||||
EndInput => Right $ mapMaybe sequence res
|
||||
_ => let char = assert_total strIndex str 0 in
|
||||
throwError $ Err {reason, line, col, char}
|
||||
Left $ Err {reason, line, col, char}
|
||||
|
|
115
lib/src/Quox/Parser.idr
Normal file
115
lib/src/Quox/Parser.idr
Normal file
|
@ -0,0 +1,115 @@
|
|||
module Quox.Parser
|
||||
|
||||
import Quox.Syntax
|
||||
import Quox.Token
|
||||
import Quox.Lexer
|
||||
|
||||
import Data.SnocVect
|
||||
import Data.SnocList
|
||||
import Text.Parser
|
||||
|
||||
%default total
|
||||
|
||||
|
||||
public export
|
||||
Vars : Nat -> Type
|
||||
Vars n = SnocVect n String
|
||||
|
||||
public export
|
||||
Grammar : Bool -> Type -> Type
|
||||
Grammar = Core.Grammar () Token
|
||||
%hide Core.Grammar
|
||||
|
||||
public export
|
||||
data Error
|
||||
= Lex (Lexer.Error)
|
||||
| Parse (List1 (ParsingError Token))
|
||||
| Leftover (List BToken)
|
||||
%hide Lexer.Error
|
||||
|
||||
|
||||
public export
|
||||
parseAll : {c : Bool} -> Grammar c a -> List BToken -> Either Error a
|
||||
parseAll grm input =
|
||||
case parse grm input of
|
||||
Right (x, []) => Right x
|
||||
Right (x, rest) => Left $ Leftover rest
|
||||
Left errs => Left $ Parse errs
|
||||
|
||||
public export
|
||||
lexParseAll : {c : Bool} -> Grammar c a -> String -> Either Error a
|
||||
lexParseAll grm = lex' >=> parseAll grm
|
||||
where lex' : String -> Either Error (List BToken)
|
||||
lex' = bimap Lex id . lex
|
||||
|
||||
|
||||
|
||||
punc : Punc -> Grammar True ()
|
||||
punc p = terminal (show p) $ \case
|
||||
P p' => if p == p' then Just () else Nothing
|
||||
_ => Nothing
|
||||
|
||||
between : Punc -> Punc -> Grammar True a -> Grammar True a
|
||||
between opener closer inner = punc opener *> inner <* punc closer
|
||||
|
||||
parens, squares, braces : Grammar True a -> Grammar True a
|
||||
parens = between LParen RParen
|
||||
squares = between LSquare RSquare
|
||||
braces = between LBrace RBrace
|
||||
|
||||
|
||||
export
|
||||
number : Grammar True Nat
|
||||
number = terminal "number" $ \case
|
||||
N Zero => Just 0
|
||||
N One => Just 1
|
||||
N (Other k) => Just k
|
||||
_ => Nothing
|
||||
|
||||
zero, one, omega : Grammar True ()
|
||||
zero = terminal "0" $ \case N Zero => Just (); _ => Nothing
|
||||
one = terminal "1" $ \case N One => Just (); _ => Nothing
|
||||
omega = terminal "ω" $ \case K Omega => Just (); _ => Nothing
|
||||
|
||||
export
|
||||
quantity : Grammar True Qty
|
||||
quantity = Zero <$ zero <|> One <$ one <|> Any <$ omega
|
||||
|
||||
|
||||
find1 : Eq a => SnocVect k a -> a -> Maybe (Var k)
|
||||
find1 [<] y = Nothing
|
||||
find1 (sx :< x) y = if x == y then Just VZ else VS <$> find1 sx y
|
||||
|
||||
find : Vars k -> Name -> Maybe (Var k)
|
||||
find vs (MakeName [<] (UN y)) = find1 vs y
|
||||
find _ _ = Nothing
|
||||
|
||||
|
||||
export
|
||||
bound : Vars k -> Grammar True (Var k)
|
||||
bound vs =
|
||||
terminal "bound variable" $ \case Name x => find1 vs x; _ => Nothing
|
||||
|
||||
export
|
||||
sname : Grammar True String
|
||||
sname = terminal "simple name" $ \case Name x => pure x; _ => Nothing
|
||||
|
||||
export
|
||||
qname : Grammar True Name
|
||||
qname = do
|
||||
parts <- sepBy1 (punc Dot) sname
|
||||
pure $ MakeName {mods = cast $ init parts, base = UN $ last parts}
|
||||
|
||||
export
|
||||
nameWith : Vars k -> Grammar True (Either (Var k) Name)
|
||||
nameWith vs = do
|
||||
y <- qname
|
||||
pure $ maybe (Right y) Left $ find vs y
|
||||
|
||||
|
||||
export
|
||||
dimension : Vars d -> Grammar True (Dim d)
|
||||
dimension vs =
|
||||
K Zero <$ zero
|
||||
<|> K One <$ one
|
||||
<|> B <$> bound vs
|
|
@ -1,6 +1,7 @@
|
|||
module Quox.Token
|
||||
|
||||
import Generics.Derive
|
||||
import Text.Lexer
|
||||
|
||||
%default total
|
||||
%language ElabReflection
|
||||
|
@ -44,3 +45,8 @@ data Token
|
|||
| K Keyword
|
||||
|
||||
%runElab derive "Token" [Generic, Meta, Eq, Ord, DecEq, Show]
|
||||
|
||||
|
||||
public export
|
||||
BToken : Type
|
||||
BToken = WithBounds Token
|
||||
|
|
|
@ -1,6 +1,6 @@
|
|||
package quox-tests
|
||||
|
||||
depends = base, contrib, elab-util, sop, quox-lib
|
||||
depends = base, contrib, elab-util, sop, snocvect, quox-lib
|
||||
|
||||
executable = quox-tests
|
||||
sourcedir = "src"
|
||||
|
|
|
@ -3,12 +3,14 @@ module Tests
|
|||
import Options
|
||||
import TAP
|
||||
import Tests.Lexer
|
||||
import Tests.Parser
|
||||
import Tests.Equal
|
||||
import System
|
||||
|
||||
|
||||
allTests = [
|
||||
Lexer.tests,
|
||||
Parser.tests,
|
||||
Equal.tests
|
||||
]
|
||||
|
||||
|
|
|
@ -52,7 +52,7 @@ parameters (label : String) (input : String)
|
|||
|
||||
rejects : Test
|
||||
rejects = testThrows label (\case LexerError _ => True; _ => False) $ delay $
|
||||
bimap LexerError (map val) $ lex {m = Either RealError} input
|
||||
bimap LexerError (map val) $ lex input
|
||||
|
||||
parameters (input : String) {default False esc : Bool}
|
||||
show' : String -> String
|
||||
|
|
77
tests/src/Tests/Parser.idr
Normal file
77
tests/src/Tests/Parser.idr
Normal file
|
@ -0,0 +1,77 @@
|
|||
module Tests.Parser
|
||||
|
||||
import Quox.Syntax
|
||||
import Quox.Parser
|
||||
import Quox.Lexer
|
||||
import Tests.Lexer
|
||||
|
||||
import Data.SnocVect
|
||||
import Text.Parser
|
||||
import TAP
|
||||
|
||||
|
||||
export
|
||||
Show tok => ToInfo (ParsingError tok) where
|
||||
toInfo (Error msg Nothing) = [("msg", msg)]
|
||||
toInfo (Error msg (Just loc)) = [("loc", show loc), ("msg", msg)]
|
||||
|
||||
|
||||
numberErrs : List1 Info -> Info
|
||||
numberErrs (head ::: []) = head
|
||||
numberErrs (head ::: tail) = go 0 (head :: tail) where
|
||||
number1 : Nat -> Info -> Info
|
||||
number1 n = map $ \(k, v) => (show n ++ k, v)
|
||||
|
||||
go : Nat -> List Info -> Info
|
||||
go k [] = []
|
||||
go k (x :: xs) = number1 k x ++ go (S k) xs
|
||||
|
||||
export
|
||||
ToInfo Parser.Error where
|
||||
toInfo (Lex err) = toInfo err
|
||||
toInfo (Parse errs) = numberErrs $ map toInfo errs
|
||||
toInfo (Leftover toks) = toInfo [("leftover", toks)]
|
||||
|
||||
|
||||
RealError = Quox.Parser.Error
|
||||
%hide Lexer.RealError
|
||||
%hide Quox.Parser.Error
|
||||
|
||||
data Error a
|
||||
= Parser RealError
|
||||
| Unexpected a a
|
||||
| ShouldFail a
|
||||
|
||||
export
|
||||
Show a => ToInfo (Error a) where
|
||||
toInfo (Parser err) = toInfo err
|
||||
toInfo (Unexpected exp got) = toInfo $
|
||||
[("expected", exp), ("received", got)]
|
||||
toInfo (ShouldFail got) = toInfo [("success", got)]
|
||||
|
||||
|
||||
parameters {c : Bool} (grm : Grammar c a) (input : String)
|
||||
parses : (Show a, Eq a) => a -> Test
|
||||
parses exp = test input $ delay $
|
||||
case lexParseAll grm input of
|
||||
Right got => if got == exp then Right ()
|
||||
else Left $ Unexpected exp got
|
||||
Left err => Left $ Parser err
|
||||
|
||||
rejects : Show a => Test
|
||||
rejects = test input $ do
|
||||
case lexParseAll grm input of
|
||||
Left err => Right ()
|
||||
Right val => Left $ ShouldFail val
|
||||
|
||||
tests = "parser" :- [
|
||||
"numbers" :- let parses = parses number in [
|
||||
parses "0" 0,
|
||||
parses "1" 1,
|
||||
parses "1000" 1000
|
||||
],
|
||||
|
||||
"bound vars" :- let parses = parses (bound [< "x", "y", "z"]) in [
|
||||
|
||||
]
|
||||
]
|
Loading…
Reference in a new issue