Extensible parsing via attributes
Equations
- Lean.Parser.registerBuiltinNodeKind k = ST.Ref.modify Lean.Parser.builtinSyntaxNodeKindSetRef fun (s : Lean.Parser.SyntaxNodeKindSet) => s.insert k
Instances For
- token: Lean.Parser.Token → Lean.Parser.ParserExtension.OLeanEntry
- kind: Lean.SyntaxNodeKind → Lean.Parser.ParserExtension.OLeanEntry
- category: Lake.Name → Lake.Name → Lean.Parser.LeadingIdentBehavior → Lean.Parser.ParserExtension.OLeanEntry
- parser: Lake.Name → Lake.Name → Nat → Lean.Parser.ParserExtension.OLeanEntry
Instances For
- token: Lean.Parser.Token → Lean.Parser.ParserExtension.Entry
- kind: Lean.SyntaxNodeKind → Lean.Parser.ParserExtension.Entry
- category: Lake.Name → Lake.Name → Lean.Parser.LeadingIdentBehavior → Lean.Parser.ParserExtension.Entry
- parser: Lake.Name → Lake.Name → Bool → Lean.Parser.Parser → Nat → Lean.Parser.ParserExtension.Entry
Instances For
Equations
- Lean.Parser.ParserExtension.instInhabitedEntry = { default := Lean.Parser.ParserExtension.Entry.token default }
Equations
- One or more equations did not get rendered due to their size.
Instances For
- tokens : Lean.Parser.TokenTable
- kinds : Lean.Parser.SyntaxNodeKindSet
- categories : Lean.Parser.ParserCategories
Instances For
Equations
- Lean.Parser.ParserExtension.instInhabitedState = { default := { tokens := default, kinds := default, categories := default } }
Equations
Instances For
Equations
- Lean.Parser.getCategory categories catName = Lean.PersistentHashMap.find? categories catName
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.addParserTokens tokenTable info = let newTokens := info.collectTokens []; List.foldlM Lean.Parser.addTokenConfig tokenTable newTokens
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Parser aliases for making ParserDescr extensible
- const: {α : Type} → α → Lean.Parser.AliasValue α
- unary: {α : Type} → (α → α) → Lean.Parser.AliasValue α
- binary: {α : Type} → (α → α → α) → Lean.Parser.AliasValue α
Instances For
Equations
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.getAlias mapRef aliasName = do let __do_lift ← ST.Ref.get mapRef pure (Lean.NameMap.find? __do_lift aliasName)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.instCoeParserParserAliasValue = { coe := Lean.Parser.AliasValue.const }
Equations
- Lean.Parser.instCoeForallParserParserAliasValue = { coe := Lean.Parser.AliasValue.unary }
Equations
- Lean.Parser.instCoeForallParserForallParserAliasValue = { coe := Lean.Parser.AliasValue.binary }
Equations
- Lean.Parser.isParserAlias aliasName = do let __do_lift ← Lean.Parser.getAlias Lean.Parser.parserAliasesRef aliasName match __do_lift with | some val => pure true | x => pure false
Instances For
Equations
- Lean.Parser.getSyntaxKindOfParserAlias? aliasName = do let __do_lift ← ST.Ref.get Lean.Parser.parserAlias2kindRef pure (__do_lift.find? aliasName)
Instances For
Equations
- Lean.Parser.ensureUnaryParserAlias aliasName = discard (Lean.Parser.getUnaryAlias Lean.Parser.parserAliasesRef aliasName)
Instances For
Equations
- Lean.Parser.ensureBinaryParserAlias aliasName = discard (Lean.Parser.getBinaryAlias Lean.Parser.parserAliasesRef aliasName)
Instances For
Equations
- Lean.Parser.ensureConstantParserAlias aliasName = discard (Lean.Parser.getConstAlias Lean.Parser.parserAliasesRef aliasName)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.compileParserDescr categories d = Lean.Parser.compileParserDescr.visit categories d
Instances For
Equations
- Lean.Parser.mkParserOfConstant categories constName = Lean.Parser.mkParserOfConstantAux constName (Lean.Parser.compileParserDescr categories)
Instances For
- postAdd : Lake.Name → Lake.Name → Bool → Lean.AttrM UnitCalled after a parser attribute is applied to a declaration. 
Instances For
Equations
- Lean.Parser.registerParserAttributeHook hook = ST.Ref.modify Lean.Parser.parserAttributeHooks fun (hooks : List Lean.Parser.ParserAttributeHook) => hook :: hooks
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.isParserCategory env catName = Lean.PersistentHashMap.contains (Lean.ScopedEnvExtension.getState Lean.Parser.parserExtension env).categories catName
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Run declName if possible and inside a quotation, or else p. The ParserInfo will always be taken from p.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.addBuiltinLeadingParser catName declName p prio = Lean.Parser.addBuiltinParser catName declName true p prio
Instances For
Equations
- Lean.Parser.addBuiltinTrailingParser catName declName p prio = Lean.Parser.addBuiltinParser catName declName false p prio
Instances For
Equations
- Lean.Parser.mkCategoryAntiquotParser kind = Lean.Parser.mkAntiquot kind.toString kind true true
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
Instances For
Equations
- Lean.Parser.mkParserState input = { stxStack := Lean.Parser.SyntaxStack.empty, lhsPrec := 0, pos := 0, cache := Lean.Parser.initCacheForInput input, errorMsg := none, recoveredErrors := #[] }
Instances For
convenience function for testing
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.declareLeadingBuiltinParser catName declName prio = Lean.Parser.declareBuiltinParser `Lean.Parser.addBuiltinLeadingParser catName declName prio
Instances For
Equations
- Lean.Parser.declareTrailingBuiltinParser catName declName prio = Lean.Parser.declareBuiltinParser `Lean.Parser.addBuiltinTrailingParser catName declName prio
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
The parsing tables for builtin parsers are "stored" in the extracted source code.
Equations
- One or more equations did not get rendered due to their size.
Instances For
A builtin parser attribute that can be extended by users.
Equations
- Lean.Parser.registerBuiltinDynamicParserAttribute attrName catName ref = Lean.registerBuiltinAttribute (Lean.Parser.mkParserAttributeImpl attrName catName ref)
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- Lean.Parser.commandParser rbp = Lean.Parser.categoryParser `command rbp
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
If the parsing stack is of the form #[.., openCommand], we process the open command, and execute p
Equations
- One or more equations did not get rendered due to their size.
Instances For
Instances For
If the parsing stack is of the form #[.., openDecl], we process the open declaration, and execute p
Equations
- Lean.Parser.withOpenDeclFn p c s = if s.stxStack.size > 0 then let stx := s.stxStack.back; Lean.Parser.withOpenDeclFnCore stx p c s else p c s
Instances For
Instances For
Helper environment extension that gives us access to built-in aliases in pure parser functions.
Result of resolving a parser name.
- category: Lake.Name → Lean.Parser.ParserResolutionReference to a category. 
- parser: Lake.Name → Bool → Lean.Parser.ParserResolutionReference to a parser declaration in the environment. A (Trailing)ParserDescrifisDescris true.
- alias: Lean.Parser.ParserAliasValue → Lean.Parser.ParserResolutionReference to a parser alias. Note that as aliases are built-in, a corresponding declaration may not be in the environment (yet). 
Instances For
Resolve the given parser name and return a list of candidates.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Resolve the given parser name and return a list of candidates.
Equations
- ctx.resolveParserName id = Lean.Parser.resolveParserNameCore ctx.env ctx.currNamespace ctx.openDecls id
Instances For
Resolve the given parser name and return a list of candidates.
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.
Instances For
Equations
- One or more equations did not get rendered due to their size.