add #![log] pragma

This commit is contained in:
rhiannon morris 2024-04-12 21:49:15 +02:00
parent f56f594839
commit 9d60f366cf
7 changed files with 164 additions and 34 deletions

View file

@ -2,6 +2,7 @@ module Quox.Log
import Quox.Loc import Quox.Loc
import Quox.Pretty import Quox.Pretty
import Quox.PrettyValExtra
import Data.So import Data.So
import Data.DPair import Data.DPair
@ -90,7 +91,7 @@ record LogLevels where
defLevel : LogLevel defLevel : LogLevel
levels : LevelMap levels : LevelMap
%name LogLevels lvls %name LogLevels lvls
%runElab derive "LogLevels" [Eq, Show] %runElab derive "LogLevels" [Eq, Show, PrettyVal]
public export public export
LevelStack : Type LevelStack : Type
@ -142,6 +143,7 @@ data PushArg =
SetDefault LogLevel SetDefault LogLevel
| SetCat LogCategory LogLevel | SetCat LogCategory LogLevel
| SetAll LogLevel | SetAll LogLevel
%runElab derive "PushArg" [Eq, Ord, Show, PrettyVal]
%name PushArg push %name PushArg push
export %inline export %inline

View file

@ -398,6 +398,10 @@ fromPDecl (PDef def) =
fromPDecl (PNs ns) = fromPDecl (PNs ns) =
maybeFail ns.fail ns.loc $ maybeFail ns.fail ns.loc $
localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls localAt NS (<+> ns.name) $ concat <$> traverse fromPDecl ns.decls
fromPDecl (PPrag prag) =
case prag of
PLogPush p _ => Log.push p $> []
PLogPop _ => Log.pop $> []
mutual mutual
export covering export covering

View file

@ -247,7 +247,7 @@ public export
reserved : List Reserved reserved : List Reserved
reserved = reserved =
[Punc1 "(", Punc1 ")", Punc1 "[", Punc1 "]", Punc1 "{", Punc1 "}", [Punc1 "(", Punc1 ")", Punc1 "[", Punc1 "]", Punc1 "{", Punc1 "}",
Punc1 ",", Punc1 ";", Punc1 "#[", Punc1 ",", Punc1 ";", Punc1 "#[", Punc1 "#![",
Sym1 "@", Sym1 "@",
Sym1 ":", Sym1 ":",
Sym "" `Or` Sym "=>", Sym "" `Or` Sym "=>",

View file

@ -626,14 +626,19 @@ term fname = lamTerm fname
export export
attr : FileName -> Grammar True PAttr attr' : FileName -> (o : String) -> (0 _ : IsReserved o) =>
attr fname = withLoc fname $ do Grammar True PAttr
resC "#[" attr' fname o = withLoc fname $ do
resC o
name <- baseName name <- baseName
args <- many $ termArg fname args <- many $ termArg fname
mustWork $ resC "]" mustWork $ resC "]"
pure $ PA name args pure $ PA name args
export %inline
attr : FileName -> Grammar True PAttr
attr fname = attr' fname "#["
export export
findDups : List PAttr -> List String findDups : List PAttr -> List String
findDups attrs = findDups attrs =
@ -658,44 +663,48 @@ attrList fname = do
noDups res $> res noDups res $> res
public export public export
data AttrMatch a = Matched a | NoMatch String | Malformed String String data AttrMatch a =
Matched a
| NoMatch String (List String)
| Malformed String String
export export
Functor AttrMatch where Functor AttrMatch where
map f (Matched x) = Matched $ f x map f (Matched x) = Matched $ f x
map f (NoMatch s) = NoMatch s map f (NoMatch s w) = NoMatch s w
map f (Malformed a e) = Malformed a e map f (Malformed a e) = Malformed a e
export export
(<|>) : AttrMatch a -> AttrMatch a -> AttrMatch a (<|>) : AttrMatch a -> AttrMatch a -> AttrMatch a
Matched x <|> _ = Matched x Matched x <|> _ = Matched x
NoMatch _ <|> y = y NoMatch {} <|> y = y
Malformed a e <|> _ = Malformed a e Malformed a e <|> _ = Malformed a e
export export
isFail : PAttr -> AttrMatch PFail isFail : PAttr -> List String -> AttrMatch PFail
isFail (PA "fail" [] _) = Matched PFailAny isFail (PA "fail" [] _) _ = Matched PFailAny
isFail (PA "fail" [Str s _] _) = Matched $ PFailMatch s isFail (PA "fail" [Str s _] _) _ = Matched $ PFailMatch s
isFail (PA "fail" _ _) = Malformed "fail" "be absent or a string literal" isFail (PA "fail" _ _) _ = Malformed "fail" "be absent or a string literal"
isFail a = NoMatch a.name isFail a w = NoMatch a.name w
export export
isMain : PAttr -> AttrMatch () isMain : PAttr -> List String -> AttrMatch ()
isMain (PA "main" [] _) = Matched () isMain (PA "main" [] _) _ = Matched ()
isMain (PA "main" _ _) = Malformed "main" "have no arguments" isMain (PA "main" _ _) _ = Malformed "main" "have no arguments"
isMain a = NoMatch a.name isMain a w = NoMatch a.name w
export export
isScheme : PAttr -> AttrMatch String isScheme : PAttr -> List String -> AttrMatch String
isScheme (PA "compile-scheme" [Str s _] _) = Matched s isScheme (PA "compile-scheme" [Str s _] _) _ = Matched s
isScheme (PA "compile-scheme" _ _) = isScheme (PA "compile-scheme" _ _) _ =
Malformed "compile-scheme" "be a string literal" Malformed "compile-scheme" "be a string literal"
isScheme a = NoMatch a.name isScheme a w = NoMatch a.name w
export export
matchAttr : String -> AttrMatch a -> Either String a matchAttr : String -> AttrMatch a -> Either String a
matchAttr _ (Matched x) = Right x matchAttr _ (Matched x) = Right x
matchAttr d (NoMatch a) = Left "unrecognised \{d} attribute \{a}" matchAttr d (NoMatch a w) = Left $ unlines
["unrecognised \{d} attribute \{a}", "expected one of: \{show w}"]
matchAttr _ (Malformed a s) = Left $ unlines matchAttr _ (Malformed a s) = Left $ unlines
["invalid \{a} attribute", "(should \{s})"] ["invalid \{a} attribute", "(should \{s})"]
@ -710,10 +719,12 @@ where
data PDefAttr = DefFail PFail | DefMain | DefScheme String data PDefAttr = DefFail PFail | DefMain | DefScheme String
isDefAttr : PAttr -> Either String PDefAttr isDefAttr : PAttr -> Either String PDefAttr
isDefAttr attr = matchAttr "definition" $ isDefAttr attr =
DefFail <$> isFail attr let defAttrs = ["fail", "main", "compile-scheme"] in
<|> DefMain <$ isMain attr matchAttr "definition" $
<|> DefScheme <$> isScheme attr DefFail <$> isFail attr defAttrs
<|> DefMain <$ isMain attr defAttrs
<|> DefScheme <$> isScheme attr defAttrs
addAttr : PDefinition -> PAttr -> Either String PDefinition addAttr : PDefinition -> PAttr -> Either String PDefinition
addAttr def attr = addAttr def attr =
@ -730,7 +741,7 @@ mkPNamespace attrs name decls = do
res <- foldlM addAttr start attrs res <- foldlM addAttr start attrs
pure $ \l => {loc_ := l} (the PNamespace res) pure $ \l => {loc_ := l} (the PNamespace res)
where where
isNsAttr = matchAttr "namespace" . isFail isNsAttr a = matchAttr "namespace" $ isFail a ["fail"]
addAttr : PNamespace -> PAttr -> Either String PNamespace addAttr : PNamespace -> PAttr -> Either String PNamespace
addAttr ns attr = pure $ {fail := !(isNsAttr attr)} ns addAttr ns attr = pure $ {fail := !(isNsAttr attr)} ns
@ -785,6 +796,48 @@ export
nsname : Grammar True Mods nsname : Grammar True Mods
nsname = do ns <- qname; pure $ ns.mods :< ns.base nsname = do ns <- qname; pure $ ns.mods :< ns.base
export
pragma : FileName -> Grammar True PPragma
pragma fname = do
a <- attr' fname "#!["
either fatalError pure $ case a.name of
"log" => logArgs a.args a.loc
_ => Left $
#"unrecognised pragma "\#{a.name}"\n"# ++
#"known pragmas: ["log"]"#
where
levelOOB : Nat -> Either String a
levelOOB n = Left $
"log level \{show n} out of bounds\n" ++
"expected number in range 0\{show maxLogLevel} inclusive"
toLevel : Nat -> Either String LogLevel
toLevel lvl = maybe (levelOOB lvl) Right $ toLogLevel lvl
unknownCat : String -> Either String a
unknownCat cat = Left $
"unknown log category \{show cat}\n" ++
"known categories: \{show $ ["all", "default"] ++ logCategories}"
toCat : String -> Either String LogCategory
toCat cat = maybe (unknownCat cat) Right $ toLogCategory cat
fromPair : PTerm -> Either String (String, Nat)
fromPair (Pair (V (MkPName [<] x) Nothing _) (Nat n _) _) = Right (x, n)
fromPair _ = Left "invalid argument to log pragma"
logCatArg : (String, Nat) -> Either String Log.PushArg
logCatArg ("default", lvl) = [|SetDefault $ toLevel lvl|]
logCatArg ("all", lvl) = [|SetAll $ toLevel lvl|]
logCatArg (cat, lvl) = [|SetCat (toCat cat) (toLevel lvl)|]
logArgs : List PTerm -> Loc -> Either String PPragma
logArgs [] _ = Left "missing arguments to log pragma"
logArgs [V "pop" Nothing _] loc = Right $ PLogPop loc
logArgs other loc = do
args <- traverse (logCatArg <=< fromPair) other
pure $ PLogPush args loc
export export
decl : FileName -> Grammar True PDecl decl : FileName -> Grammar True PDecl
@ -806,7 +859,9 @@ declBody fname attrs =
[|PDef $ definition fname attrs|] <|> [|PNs $ namespace_ fname attrs|] [|PDef $ definition fname attrs|] <|> [|PNs $ namespace_ fname attrs|]
-- decl : FileName -> Grammar True PDecl -- decl : FileName -> Grammar True PDecl
decl fname = attrList fname >>= declBody fname decl fname =
(attrList fname >>= declBody fname)
<|> PPrag <$> pragma fname
export export
load : FileName -> Grammar True PTopLevel load : FileName -> Grammar True PTopLevel

View file

@ -4,6 +4,7 @@ import public Quox.Loc
import public Quox.Syntax import public Quox.Syntax
import public Quox.Definition import public Quox.Definition
import Quox.PrettyValExtra import Quox.PrettyValExtra
import public Quox.Log
import Derive.Prelude import Derive.Prelude
%hide TT.Name %hide TT.Name
@ -184,6 +185,18 @@ record PDefinition where
export Located PDefinition where def.loc = def.loc_ export Located PDefinition where def.loc = def.loc_
public export
data PPragma =
PLogPush (List Log.PushArg) Loc
| PLogPop Loc
%name PPragma prag
%runElab derive "PPragma" [Eq, Ord, Show, PrettyVal]
export
Located PPragma where
(PLogPush _ loc).loc = loc
(PLogPop loc).loc = loc
mutual mutual
public export public export
record PNamespace where record PNamespace where
@ -196,8 +209,9 @@ mutual
public export public export
data PDecl = data PDecl =
PDef PDefinition PDef PDefinition
| PNs PNamespace | PNs PNamespace
| PPrag PPragma
%name PDecl decl %name PDecl decl
%runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show, PrettyVal] %runElab deriveMutual ["PNamespace", "PDecl"] [Eq, Ord, Show, PrettyVal]
@ -205,8 +219,9 @@ export Located PNamespace where ns.loc = ns.loc_
export export
Located PDecl where Located PDecl where
(PDef d).loc = d.loc (PDef d).loc = d.loc
(PNs ns).loc = ns.loc (PNs ns).loc = ns.loc
(PPrag prag).loc = prag.loc
public export public export
data PTopLevel = PD PDecl | PLoad String Loc data PTopLevel = PD PDecl | PLoad String Loc

View file

@ -1,5 +1,6 @@
module Quox.PrettyValExtra module Quox.PrettyValExtra
import Data.DPair
import Derive.Prelude import Derive.Prelude
import public Text.Show.Value import public Text.Show.Value
import public Text.Show.PrettyVal import public Text.Show.PrettyVal
@ -8,3 +9,12 @@ import public Text.Show.PrettyVal.Derive
%language ElabReflection %language ElabReflection
%runElab derive "SnocList" [PrettyVal] %runElab derive "SnocList" [PrettyVal]
export %inline
PrettyVal a => PrettyVal (Subset a p) where
prettyVal (Element x _) = Con "Element" [prettyVal x, Con "_" []]
export %inline
(forall x. PrettyVal (p x)) => PrettyVal (Exists p) where
prettyVal (Evidence _ p) = Con "Evidence" [Con "_" [], prettyVal p]

View file

@ -531,7 +531,51 @@ tests = "parser" :- [
parseMatch input #" load "a.quox"; def b = a.b "# parseMatch input #" load "a.quox"; def b = a.b "#
`([PLoad "a.quox" _, `([PLoad "a.quox" _,
PD (PDef $ MkPDef (PQ Any _) "b" PD (PDef $ MkPDef (PQ Any _) "b"
(PConcrete Nothing (V (MakePName [< "a"] "b") Nothing _)) (PConcrete Nothing (V (MkPName [< "a"] "b") Nothing _))
PSucceed False Nothing _)]) PSucceed False Nothing _)]),
parseMatch input #" #[main] postulate hi : String "#
`([PD (PDef $ MkPDef (PQ Any _) "hi"
(PPostulate (STRING _))
PSucceed True Nothing _)]),
parseMatch input #" #[compile-scheme "hi"] postulate hi : String "#
`([PD (PDef $ MkPDef (PQ Any _) "hi"
(PPostulate (STRING _))
PSucceed False (Just "hi") _)]),
parseMatch input #" #[main] #[compile-scheme "hi"] postulate hi : String "#
`([PD (PDef $ MkPDef (PQ Any _) "hi"
(PPostulate (STRING _))
PSucceed True (Just "hi") _)]),
parseMatch input #" #[fail] def hi = "hi!!!! uwu" "#
`([PD (PDef $ MkPDef (PQ Any _) "hi"
(PConcrete Nothing (Str "hi!!!! uwu" _))
PFailAny False Nothing _)]),
parseMatch input #" #[fail "type"] def hi = "hi!!!! uwu" "#
`([PD (PDef $ MkPDef (PQ Any _) "hi"
(PConcrete Nothing (Str "hi!!!! uwu" _))
(PFailMatch "type") False Nothing _)]),
parseMatch input #" #[fail] namespace ns { } "#
`([PD (PNs $ MkPNamespace [< "ns"] [] PFailAny _)]),
parseFails input #" #[fail 69] namespace ns { } "#,
parseFails input "#[main]",
parseFails input "#[main] namespace a { } ",
parseFails input #" #[not-an-attr] postulate hi : String "#,
parseFails input #" #[log pop] postulate hi : String "#,
parseMatch input #" #![log pop] "#
`([PD (PPrag (PLogPop _))]),
parseMatch input #" #![log (all, 5)] "#
`([PD (PPrag (PLogPush [SetAll (Element 5 _)] _))]),
parseMatch input #" #![log (default, 69)] "#
`([PD (PPrag (PLogPush [SetDefault (Element 69 _)] _))]),
parseMatch input #" #![log (whnf, 100)] "#
`([PD (PPrag (PLogPush [SetCat (Element "whnf" _) (Element 100 _)] _))]),
parseMatch input #" #![log (all, 5) (default, 69) (whnf, 100)] "#
`([PD (PPrag (PLogPush
[SetAll (Element 5 _), SetDefault (Element 69 _),
SetCat (Element "whnf" _) (Element 100 _)] _))]),
parseFails input #" #![log] "#,
parseFails input #" #![log (non-category, 5)] "#,
parseFails input #" #![log (whnf, 50000000)] "#,
parseFails input #" #![log [0.★⁵]] "#,
parseFails input #" #![main] "#
] ]
] ]