add @[fail] modifier to declarations
- `@[fail] def foo = ...` succeeds if `foo` has some error. - `@[fail "scope"] def foo = ...` succeeds if `foo` has some error containing the word "scope" somewhere - `@[fail] namespace foo { }` works too and the error must be anywhere in the namespace
This commit is contained in:
parent
ea674503c0
commit
d4cfbd4045
6 changed files with 127 additions and 41 deletions
|
@ -3,6 +3,7 @@ module Quox.Parser.FromParser
|
||||||
|
|
||||||
import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
|
import public Quox.Parser.FromParser.Error as Quox.Parser.FromParser
|
||||||
|
|
||||||
|
import Quox.Pretty
|
||||||
import Quox.Parser.Syntax
|
import Quox.Parser.Syntax
|
||||||
import Quox.Parser.Parser
|
import Quox.Parser.Parser
|
||||||
import public Quox.Parser.LoadFile
|
import public Quox.Parser.LoadFile
|
||||||
|
@ -42,6 +43,35 @@ FromParserIO : List (Type -> Type)
|
||||||
FromParserIO = LoadFile :: FromParserPure
|
FromParserIO = LoadFile :: FromParserPure
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
fromParserPure : NameSuf -> Definitions ->
|
||||||
|
Eff FromParserPure a ->
|
||||||
|
Either Error (a, NameSuf, Definitions)
|
||||||
|
fromParserPure suf defs act = runSTErr $ do
|
||||||
|
suf <- liftST $ newSTRef suf
|
||||||
|
defs <- liftST $ newSTRef defs
|
||||||
|
res <- runEff act $ with Union.(::)
|
||||||
|
[handleExcept (\e => stLeft e),
|
||||||
|
handleStateSTRef defs,
|
||||||
|
handleStateSTRef !(liftST $ newSTRef [<]),
|
||||||
|
handleStateSTRef suf]
|
||||||
|
pure (res, !(liftST $ readSTRef suf), !(liftST $ readSTRef defs))
|
||||||
|
|
||||||
|
|
||||||
|
export covering
|
||||||
|
fromParserIO : (MonadRec io, HasIO io) =>
|
||||||
|
IncludePath -> IORef SeenSet ->
|
||||||
|
IORef NameSuf -> IORef Definitions ->
|
||||||
|
Eff FromParserIO a -> io (Either Error a)
|
||||||
|
fromParserIO inc seen suf defs act = liftIO $ fromIOErr $ do
|
||||||
|
runEff act $ with Union.(::)
|
||||||
|
[handleLoadFileIOE LoadError seen inc,
|
||||||
|
handleExcept (\e => ioLeft e),
|
||||||
|
handleStateIORef defs,
|
||||||
|
handleStateIORef !(newIORef [<]),
|
||||||
|
handleStateIORef suf]
|
||||||
|
|
||||||
|
|
||||||
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
|
parameters {auto _ : Functor m} (b : Var n -> m a) (f : PName -> m a)
|
||||||
(xs : Context' PatVar n)
|
(xs : Context' PatVar n)
|
||||||
private
|
private
|
||||||
|
@ -318,12 +348,43 @@ fromPDef (MkPDef qty pname ptype pterm defLoc) = do
|
||||||
res <- liftTC $ inferC empty sqty elim
|
res <- liftTC $ inferC empty sqty elim
|
||||||
addDef name gqty res.type term defLoc
|
addDef name gqty res.type term defLoc
|
||||||
|
|
||||||
|
|
||||||
|
public export
|
||||||
|
data HasFail = NoFail | AnyFail | FailWith String
|
||||||
|
|
||||||
|
export
|
||||||
|
hasFail : List PDeclMod -> HasFail
|
||||||
|
hasFail [] = NoFail
|
||||||
|
hasFail (PFail str _ :: _) = maybe AnyFail FailWith str
|
||||||
|
|
||||||
export covering
|
export covering
|
||||||
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
fromPDecl : PDecl -> Eff FromParserPure (List NDefinition)
|
||||||
fromPDecl (PDef def) = singleton <$> fromPDef def
|
|
||||||
fromPDecl (PNs ns) =
|
export covering
|
||||||
|
fromPDeclBody : PDeclBody -> Eff FromParserPure (List NDefinition)
|
||||||
|
fromPDeclBody (PDef def) = singleton <$> fromPDef def
|
||||||
|
fromPDeclBody (PNs ns) =
|
||||||
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
|
||||||
|
|
||||||
|
export covering
|
||||||
|
expectFail : PDeclBody -> Eff FromParserPure Error
|
||||||
|
expectFail body =
|
||||||
|
case fromParserPure !(getAt GEN) !(getAt DEFS) $ fromPDeclBody body of
|
||||||
|
Left err => pure err
|
||||||
|
Right _ => throw $ ExpectedFail body.loc
|
||||||
|
|
||||||
|
|
||||||
|
fromPDecl (MkPDecl mods decl loc) = case hasFail mods of
|
||||||
|
NoFail => fromPDeclBody decl
|
||||||
|
AnyFail => expectFail decl $> []
|
||||||
|
FailWith str => do
|
||||||
|
err <- expectFail decl
|
||||||
|
let msg = runPretty $ prettyError False err {opts = Opts 10_000} -- w/e
|
||||||
|
if str `isInfixOf` renderInfinite msg
|
||||||
|
then pure []
|
||||||
|
else throw $ WrongFail str err loc
|
||||||
|
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
export covering
|
export covering
|
||||||
loadProcessFile : Loc -> String -> Eff FromParserIO (List NDefinition)
|
loadProcessFile : Loc -> String -> Eff FromParserIO (List NDefinition)
|
||||||
|
@ -339,32 +400,3 @@ mutual
|
||||||
fromPTopLevel : PTopLevel -> Eff FromParserIO (List NDefinition)
|
fromPTopLevel : PTopLevel -> Eff FromParserIO (List NDefinition)
|
||||||
fromPTopLevel (PD decl) = lift $ fromPDecl decl
|
fromPTopLevel (PD decl) = lift $ fromPDecl decl
|
||||||
fromPTopLevel (PLoad file loc) = loadProcessFile loc file
|
fromPTopLevel (PLoad file loc) = loadProcessFile loc file
|
||||||
|
|
||||||
|
|
||||||
export
|
|
||||||
fromParserPure : NameSuf -> Definitions ->
|
|
||||||
Eff FromParserPure a ->
|
|
||||||
Either Error (a, NameSuf, Definitions)
|
|
||||||
fromParserPure suf defs act = runSTErr $ do
|
|
||||||
suf <- liftST $ newSTRef suf
|
|
||||||
defs <- liftST $ newSTRef defs
|
|
||||||
res <- runEff act $ with Union.(::)
|
|
||||||
[handleExcept (\e => stLeft e),
|
|
||||||
handleStateSTRef defs,
|
|
||||||
handleStateSTRef !(liftST $ newSTRef [<]),
|
|
||||||
handleStateSTRef suf]
|
|
||||||
pure (res, !(liftST $ readSTRef suf), !(liftST $ readSTRef defs))
|
|
||||||
|
|
||||||
|
|
||||||
export covering
|
|
||||||
fromParserIO : (MonadRec io, HasIO io) =>
|
|
||||||
IncludePath -> IORef SeenSet ->
|
|
||||||
IORef NameSuf -> IORef Definitions ->
|
|
||||||
Eff FromParserIO a -> io (Either Error a)
|
|
||||||
fromParserIO inc seen suf defs act = liftIO $ fromIOErr $ do
|
|
||||||
runEff act $ with Union.(::)
|
|
||||||
[handleLoadFileIOE LoadError seen inc,
|
|
||||||
handleExcept (\e => ioLeft e),
|
|
||||||
handleStateIORef defs,
|
|
||||||
handleStateIORef !(newIORef [<]),
|
|
||||||
handleStateIORef suf]
|
|
||||||
|
|
|
@ -30,6 +30,8 @@ data Error =
|
||||||
| DisplacedBoundVar Loc PName
|
| DisplacedBoundVar Loc PName
|
||||||
| WrapTypeError TypeError
|
| WrapTypeError TypeError
|
||||||
| LoadError Loc FilePath FileError
|
| LoadError Loc FilePath FileError
|
||||||
|
| ExpectedFail Loc
|
||||||
|
| WrongFail String Error Loc
|
||||||
| WrapParseError String ParseError
|
| WrapParseError String ParseError
|
||||||
|
|
||||||
|
|
||||||
|
@ -106,5 +108,13 @@ parameters (showContext : Bool)
|
||||||
"couldn't load file" <++> text file,
|
"couldn't load file" <++> text file,
|
||||||
text $ show err]
|
text $ show err]
|
||||||
|
|
||||||
|
prettyError (ExpectedFail loc) = pure $
|
||||||
|
sep [!(prettyLoc loc), "expected error"]
|
||||||
|
|
||||||
|
prettyError (WrongFail str err loc) = pure $
|
||||||
|
sep [!(prettyLoc loc),
|
||||||
|
"wrong error, expected to match", !(hl Tag $ text "\"\{str}\""),
|
||||||
|
"but got", !(prettyError err)]
|
||||||
|
|
||||||
prettyError (WrapParseError file err) =
|
prettyError (WrapParseError file err) =
|
||||||
prettyParseError file err
|
prettyParseError file err
|
||||||
|
|
|
@ -190,7 +190,7 @@ reserved : List Reserved
|
||||||
reserved =
|
reserved =
|
||||||
[Punc1 '(', Punc1 ')', Punc1 '[', Punc1 ']', Punc1 '{', Punc1 '}',
|
[Punc1 '(', Punc1 ')', Punc1 '[', Punc1 ']', Punc1 '{', Punc1 '}',
|
||||||
Punc1 ',', Punc1 ';',
|
Punc1 ',', Punc1 ';',
|
||||||
Sym1 "@",
|
Sym1 "@[", Sym1 "@",
|
||||||
Sym1 ":",
|
Sym1 ":",
|
||||||
Sym "⇒" `Or` Sym "=>",
|
Sym "⇒" `Or` Sym "=>",
|
||||||
Sym "→" `Or` Sym "->",
|
Sym "→" `Or` Sym "->",
|
||||||
|
|
|
@ -149,6 +149,12 @@ export
|
||||||
qty : FileName -> Grammar True PQty
|
qty : FileName -> Grammar True PQty
|
||||||
qty fname = withLoc fname [|PQ qtyVal|]
|
qty fname = withLoc fname [|PQ qtyVal|]
|
||||||
|
|
||||||
|
export
|
||||||
|
exactName : String -> Grammar True ()
|
||||||
|
exactName name = terminal "expected '\{name}'" $ \case
|
||||||
|
Name (MakePName [<] x) => guard $ x == name
|
||||||
|
_ => Nothing
|
||||||
|
|
||||||
|
|
||||||
||| pattern var (unqualified name or _)
|
||| pattern var (unqualified name or _)
|
||||||
export
|
export
|
||||||
|
@ -576,6 +582,15 @@ term fname = lamTerm fname
|
||||||
<|> sigmaTerm fname
|
<|> sigmaTerm fname
|
||||||
|
|
||||||
|
|
||||||
|
export
|
||||||
|
pragma : Grammar True a -> Grammar True a
|
||||||
|
pragma p = resC "@[" *> p <* mustWork (resC "]")
|
||||||
|
|
||||||
|
export
|
||||||
|
declMod : FileName -> Grammar True PDeclMod
|
||||||
|
declMod fname = withLoc fname $ pragma $
|
||||||
|
exactName "fail" *> [|PFail $ optional strLit|]
|
||||||
|
|
||||||
export
|
export
|
||||||
decl : FileName -> Grammar True PDecl
|
decl : FileName -> Grammar True PDecl
|
||||||
|
|
||||||
|
@ -610,7 +625,11 @@ where
|
||||||
nsInner = [] <$ resC "}"
|
nsInner = [] <$ resC "}"
|
||||||
<|> [|(assert_total decl fname <* commit) :: assert_total nsInner|]
|
<|> [|(assert_total decl fname <* commit) :: assert_total nsInner|]
|
||||||
|
|
||||||
decl fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
export
|
||||||
|
declBody : FileName -> Grammar True PDeclBody
|
||||||
|
declBody fname = [|PDef $ definition fname|] <|> [|PNs $ namespace_ fname|]
|
||||||
|
|
||||||
|
decl fname = withLoc fname [|MkPDecl (many $ declMod fname) (declBody fname)|]
|
||||||
|
|
||||||
export
|
export
|
||||||
load : FileName -> Grammar True PTopLevel
|
load : FileName -> Grammar True PTopLevel
|
||||||
|
|
|
@ -153,6 +153,12 @@ record PDefinition where
|
||||||
|
|
||||||
export Located PDefinition where def.loc = def.loc_
|
export Located PDefinition where def.loc = def.loc_
|
||||||
|
|
||||||
|
public export
|
||||||
|
data PDeclMod =
|
||||||
|
PFail (Maybe String) Loc
|
||||||
|
%name PDeclMod mod
|
||||||
|
%runElab derive "PDeclMod" [Eq, Ord, Show]
|
||||||
|
|
||||||
mutual
|
mutual
|
||||||
public export
|
public export
|
||||||
record PNamespace where
|
record PNamespace where
|
||||||
|
@ -163,16 +169,25 @@ mutual
|
||||||
%name PNamespace ns
|
%name PNamespace ns
|
||||||
|
|
||||||
public export
|
public export
|
||||||
data PDecl =
|
record PDecl where
|
||||||
|
constructor MkPDecl
|
||||||
|
mods : List PDeclMod
|
||||||
|
decl : PDeclBody
|
||||||
|
loc_ : Loc
|
||||||
|
|
||||||
|
public export
|
||||||
|
data PDeclBody =
|
||||||
PDef PDefinition
|
PDef PDefinition
|
||||||
| PNs PNamespace
|
| PNs PNamespace
|
||||||
%name PDecl decl
|
%name PDeclBody decl
|
||||||
%runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show]
|
%runElab deriveMutual ["PNamespace", "PDecl", "PDeclBody"] [Eq, Ord, Show]
|
||||||
|
|
||||||
export Located PNamespace where ns.loc = ns.loc_
|
export Located PNamespace where ns.loc = ns.loc_
|
||||||
|
|
||||||
|
export Located PDecl where decl.loc = decl.loc_
|
||||||
|
|
||||||
export
|
export
|
||||||
Located PDecl where
|
Located PDeclBody where
|
||||||
(PDef def).loc = def.loc
|
(PDef def).loc = def.loc
|
||||||
(PNs ns).loc = ns.loc
|
(PNs ns).loc = ns.loc
|
||||||
|
|
||||||
|
|
|
@ -5,6 +5,7 @@ module Text.PrettyPrint.Bernardy.Core.Decorate
|
||||||
|
|
||||||
import public Text.PrettyPrint.Bernardy.Core
|
import public Text.PrettyPrint.Bernardy.Core
|
||||||
import Data.DPair
|
import Data.DPair
|
||||||
|
import Data.String
|
||||||
|
|
||||||
|
|
||||||
public export
|
public export
|
||||||
|
@ -43,3 +44,12 @@ decorateLayout h l@(MkLayout content stats) =
|
||||||
export
|
export
|
||||||
decorate : {opts : _} -> Highlight -> Doc opts -> Doc opts
|
decorate : {opts : _} -> Highlight -> Doc opts -> Doc opts
|
||||||
decorate h doc = doc >>= \l => pure (decorateLayout h l)
|
decorate h doc = doc >>= \l => pure (decorateLayout h l)
|
||||||
|
|
||||||
|
|
||||||
|
||| render a doc with no line breaks at all
|
||||||
|
export
|
||||||
|
renderInfinite : Doc opts -> String
|
||||||
|
renderInfinite (MkDoc (MkLayout content _) _) = unwords content where
|
||||||
|
unwords : SnocList String -> String
|
||||||
|
unwords [<] = ""
|
||||||
|
unwords (xs :< x) = foldMap (++ " ") xs ++ x
|
||||||
|
|
Loading…
Reference in a new issue