diff --git a/src/Init/Data/Array/Subarray.lean b/src/Init/Data/Array/Subarray.lean index 5dc905b81052..6dc9525ac5bd 100644 --- a/src/Init/Data/Array/Subarray.lean +++ b/src/Init/Data/Array/Subarray.lean @@ -104,6 +104,28 @@ def any {α : Type u} (p : α → Bool) (as : Subarray α) : Bool := def all {α : Type u} (p : α → Bool) (as : Subarray α) : Bool := Id.run <| as.allM p +@[inline] +def findSomeRevM? {α : Type u} {β : Type v} {m : Type v → Type w} [Monad m] (as : Subarray α) (f : α → m (Option β)) : m (Option β) := + let rec @[specialize] find : (i : Nat) → i ≤ as.size → m (Option β) + | 0, _ => pure none + | i+1, h => do + have : i < as.size := Nat.lt_of_lt_of_le (Nat.lt_succ_self _) h + let r ← f as[i] + match r with + | some _ => pure r + | none => + have : i ≤ as.size := Nat.le_of_lt this + find i this + find as.size (Nat.le_refl _) + +@[inline] +def findRevM? {α : Type} {m : Type → Type w} [Monad m] (as : Subarray α) (p : α → m Bool) : m (Option α) := + as.findSomeRevM? fun a => return if (← p a) then some a else none + +@[inline] +def findRev? {α : Type} (as : Subarray α) (p : α → Bool) : Option α := + Id.run <| as.findRevM? p + end Subarray namespace Array diff --git a/src/Lean/Data/OpenDecl.lean b/src/Lean/Data/OpenDecl.lean index fad2a525648f..f6c0d0b262cf 100644 --- a/src/Lean/Data/OpenDecl.lean +++ b/src/Lean/Data/OpenDecl.lean @@ -11,6 +11,7 @@ namespace Lean inductive OpenDecl where | simple (ns : Name) (except : List Name) | explicit (id : Name) (declName : Name) + deriving BEq namespace OpenDecl instance : Inhabited OpenDecl := ⟨simple Name.anonymous []⟩ diff --git a/src/Lean/Data/Options.lean b/src/Lean/Data/Options.lean index 9e16003d2a43..dbc41f4a2c9b 100644 --- a/src/Lean/Data/Options.lean +++ b/src/Lean/Data/Options.lean @@ -16,6 +16,7 @@ instance : Inhabited Options where default := {} instance : ToString Options := inferInstanceAs (ToString KVMap) instance : ForIn m Options (Name × DataValue) := inferInstanceAs (ForIn _ KVMap _) +instance : BEq Options := inferInstanceAs (BEq KVMap) structure OptionDecl where declName : Name := by exact decl_name% diff --git a/src/Lean/Elab/BuiltinNotation.lean b/src/Lean/Elab/BuiltinNotation.lean index 8d3cf8b44626..96f7b903d391 100644 --- a/src/Lean/Elab/BuiltinNotation.lean +++ b/src/Lean/Elab/BuiltinNotation.lean @@ -89,15 +89,20 @@ private def elabParserMacroAux (prec e : Term) (withAnonymousAntiquot : Bool) : match extractMacroScopes declName with | { name := .str _ s, .. } => let kind := quote declName - let s := quote s - ``(withAntiquot (mkAntiquot $s $kind $(quote withAnonymousAntiquot)) (leadingNode $kind $prec $e)) + let mut p ← ``(withAntiquot + (mkAntiquot $(quote s) $kind $(quote withAnonymousAntiquot)) + (leadingNode $kind $prec $e)) + -- cache only unparameterized parsers + if (← getLCtx).all (·.isAuxDecl) then + p ← ``(withCache $kind $p) + return p | _ => throwError "invalid `leading_parser` macro, unexpected declaration name" @[builtin_term_elab «leading_parser»] def elabLeadingParserMacro : TermElab := - adaptExpander fun stx => match stx with - | `(leading_parser $[: $prec?]? $[(withAnonymousAntiquot := $anon?)]? $e) => - elabParserMacroAux (prec?.getD (quote Parser.maxPrec)) e (anon?.all (·.raw.isOfKind ``Parser.Term.trueVal)) - | _ => throwUnsupportedSyntax + adaptExpander fun + | `(leading_parser $[: $prec?]? $[(withAnonymousAntiquot := $anon?)]? $e) => + elabParserMacroAux (prec?.getD (quote Parser.maxPrec)) e (anon?.all (·.raw.isOfKind ``Parser.Term.trueVal)) + | _ => throwUnsupportedSyntax private def elabTParserMacroAux (prec lhsPrec e : Term) : TermElabM Syntax := do let declName? ← getDeclName? diff --git a/src/Lean/Parser/Basic.lean b/src/Lean/Parser/Basic.lean index 68568addcc93..4144d194c97b 100644 --- a/src/Lean/Parser/Basic.lean +++ b/src/Lean/Parser/Basic.lean @@ -3,9 +3,7 @@ Copyright (c) 2019 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Sebastian Ullrich -/ -import Lean.Data.Trie -import Lean.Syntax -import Lean.Message +import Lean.Parser.Types /-! # Basic Lean parser infrastructure @@ -58,272 +56,7 @@ Error recovery is left to the designer of the specific language; for example, Le tokens until the next command keyword on error. -/ -namespace Lean - -namespace Parser - -abbrev mkAtom (info : SourceInfo) (val : String) : Syntax := - Syntax.atom info val - -abbrev mkIdent (info : SourceInfo) (rawVal : Substring) (val : Name) : Syntax := - Syntax.ident info rawVal val [] - -/-- Return character after position `pos` -/ -def getNext (input : String) (pos : String.Pos) : Char := - input.get (input.next pos) - -/-- Maximal (and function application) precedence. - In the standard lean language, no parser has precedence higher than `maxPrec`. - - Note that nothing prevents users from using a higher precedence, but we strongly - discourage them from doing it. -/ -def maxPrec : Nat := eval_prec max -def argPrec : Nat := eval_prec arg -def leadPrec : Nat := eval_prec lead -def minPrec : Nat := eval_prec min - -abbrev Token := String - -structure TokenCacheEntry where - startPos : String.Pos := 0 - stopPos : String.Pos := 0 - token : Syntax := Syntax.missing - -structure ParserCache where - tokenCache : TokenCacheEntry - -def initCacheForInput (input : String) : ParserCache := { - tokenCache := { startPos := input.endPos + ' ' /- make sure it is not a valid position -/ } -} - -abbrev TokenTable := Trie Token - -abbrev SyntaxNodeKindSet := PersistentHashMap SyntaxNodeKind Unit - -def SyntaxNodeKindSet.insert (s : SyntaxNodeKindSet) (k : SyntaxNodeKind) : SyntaxNodeKindSet := - PersistentHashMap.insert s k () - -/-- - Input string and related data. Recall that the `FileMap` is a helper structure for mapping - `String.Pos` in the input string to line/column information. -/ -structure InputContext where - input : String - fileName : String - fileMap : FileMap - deriving Inhabited - -/-- Input context derived from elaboration of previous commands. -/ -structure ParserModuleContext where - env : Environment - options : Options - -- for name lookup - currNamespace : Name := .anonymous - openDecls : List OpenDecl := [] - -structure ParserContext extends InputContext, ParserModuleContext where - prec : Nat - tokens : TokenTable - -- used for bootstrapping only - quotDepth : Nat := 0 - suppressInsideQuot : Bool := false - savedPos? : Option String.Pos := none - forbiddenTk? : Option Token := none - -structure Error where - unexpected : String := "" - expected : List String := [] - deriving Inhabited, BEq - -namespace Error - -private def expectedToString : List String → String - | [] => "" - | [e] => e - | [e1, e2] => e1 ++ " or " ++ e2 - | e::es => e ++ ", " ++ expectedToString es - -protected def toString (e : Error) : String := - let unexpected := if e.unexpected == "" then [] else [e.unexpected] - let expected := if e.expected == [] then [] else - let expected := e.expected.toArray.qsort (fun e e' => e < e') - let expected := expected.toList.eraseReps - ["expected " ++ expectedToString expected] - "; ".intercalate $ unexpected ++ expected - -instance : ToString Error where - toString := Error.toString - -def merge (e₁ e₂ : Error) : Error := - match e₂ with - | { unexpected := u, .. } => { unexpected := if u == "" then e₁.unexpected else u, expected := e₁.expected ++ e₂.expected } - -end Error - -structure ParserState where - stxStack : Array Syntax := #[] - /-- - Set to the precedence of the preceding (not surrounding) parser by `runLongestMatchParser` - for the use of `checkLhsPrec` in trailing parsers. - Note that with chaining, the preceding parser can be another trailing parser: - in `1 * 2 + 3`, the preceding parser is '*' when '+' is executed. -/ - lhsPrec : Nat := 0 - pos : String.Pos := 0 - cache : ParserCache - errorMsg : Option Error := none - -namespace ParserState - -def hasError (s : ParserState) : Bool := - s.errorMsg != none - -def stackSize (s : ParserState) : Nat := - s.stxStack.size - -def restore (s : ParserState) (iniStackSz : Nat) (iniPos : String.Pos) : ParserState := - { s with stxStack := s.stxStack.shrink iniStackSz, errorMsg := none, pos := iniPos } - -def setPos (s : ParserState) (pos : String.Pos) : ParserState := - { s with pos := pos } - -def setCache (s : ParserState) (cache : ParserCache) : ParserState := - { s with cache := cache } - -def pushSyntax (s : ParserState) (n : Syntax) : ParserState := - { s with stxStack := s.stxStack.push n } - -def popSyntax (s : ParserState) : ParserState := - { s with stxStack := s.stxStack.pop } - -def shrinkStack (s : ParserState) (iniStackSz : Nat) : ParserState := - { s with stxStack := s.stxStack.shrink iniStackSz } - -def next (s : ParserState) (input : String) (pos : String.Pos) : ParserState := - { s with pos := input.next pos } - -def next' (s : ParserState) (input : String) (pos : String.Pos) (h : ¬ input.atEnd pos): ParserState := - { s with pos := input.next' pos h } - -def toErrorMsg (ctx : ParserContext) (s : ParserState) : String := - match s.errorMsg with - | none => "" - | some msg => - let pos := ctx.fileMap.toPosition s.pos - mkErrorStringWithPos ctx.fileName pos (toString msg) - -def mkNode (s : ParserState) (k : SyntaxNodeKind) (iniStackSz : Nat) : ParserState := - match s with - | ⟨stack, lhsPrec, pos, cache, err⟩ => - if err != none && stack.size == iniStackSz then - -- If there is an error but there are no new nodes on the stack, use `missing` instead. - -- Thus we ensure the property that an syntax tree contains (at least) one `missing` node - -- if (and only if) there was a parse error. - -- We should not create an actual node of kind `k` in this case because it would mean we - -- choose an "arbitrary" node (in practice the last one) in an alternative of the form - -- `node k1 p1 <|> ... <|> node kn pn` when all parsers fail. With the code below we - -- instead return a less misleading single `missing` node without randomly selecting any `ki`. - let stack := stack.push Syntax.missing - ⟨stack, lhsPrec, pos, cache, err⟩ - else - let newNode := Syntax.node SourceInfo.none k (stack.extract iniStackSz stack.size) - let stack := stack.shrink iniStackSz - let stack := stack.push newNode - ⟨stack, lhsPrec, pos, cache, err⟩ - -def mkTrailingNode (s : ParserState) (k : SyntaxNodeKind) (iniStackSz : Nat) : ParserState := - match s with - | ⟨stack, lhsPrec, pos, cache, err⟩ => - let newNode := Syntax.node SourceInfo.none k (stack.extract (iniStackSz - 1) stack.size) - let stack := stack.shrink (iniStackSz - 1) - let stack := stack.push newNode - ⟨stack, lhsPrec, pos, cache, err⟩ - -def setError (s : ParserState) (msg : String) : ParserState := - match s with - | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨stack, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ - -def mkError (s : ParserState) (msg : String) : ParserState := - match s with - | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ - -def mkUnexpectedError (s : ParserState) (msg : String) (expected : List String := []) : ParserState := - match s with - | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { unexpected := msg, expected := expected }⟩ - -def mkEOIError (s : ParserState) (expected : List String := []) : ParserState := - s.mkUnexpectedError "unexpected end of input" expected - -def mkErrorAt (s : ParserState) (msg : String) (pos : String.Pos) (initStackSz? : Option Nat := none) : ParserState := - match s, initStackSz? with - | ⟨stack, lhsPrec, _, cache, _⟩, none => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ - | ⟨stack, lhsPrec, _, cache, _⟩, some sz => ⟨stack.shrink sz |>.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ - -def mkErrorsAt (s : ParserState) (ex : List String) (pos : String.Pos) (initStackSz? : Option Nat := none) : ParserState := - match s, initStackSz? with - | ⟨stack, lhsPrec, _, cache, _⟩, none => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := ex }⟩ - | ⟨stack, lhsPrec, _, cache, _⟩, some sz => ⟨stack.shrink sz |>.push Syntax.missing, lhsPrec, pos, cache, some { expected := ex }⟩ - -def mkUnexpectedErrorAt (s : ParserState) (msg : String) (pos : String.Pos) : ParserState := - match s with - | ⟨stack, lhsPrec, _, cache, _⟩ => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { unexpected := msg }⟩ - -end ParserState - -def ParserFn := ParserContext → ParserState → ParserState - -instance : Inhabited ParserFn where - default := fun _ s => s - -inductive FirstTokens where - | epsilon : FirstTokens - | unknown : FirstTokens - | tokens : List Token → FirstTokens - | optTokens : List Token → FirstTokens - deriving Inhabited - -namespace FirstTokens - -def seq : FirstTokens → FirstTokens → FirstTokens - | epsilon, tks => tks - | optTokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) - | optTokens s₁, tokens s₂ => tokens (s₁ ++ s₂) - | tks, _ => tks - -def toOptional : FirstTokens → FirstTokens - | tokens tks => optTokens tks - | tks => tks - -def merge : FirstTokens → FirstTokens → FirstTokens - | epsilon, tks => toOptional tks - | tks, epsilon => toOptional tks - | tokens s₁, tokens s₂ => tokens (s₁ ++ s₂) - | optTokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) - | tokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) - | optTokens s₁, tokens s₂ => optTokens (s₁ ++ s₂) - | _, _ => unknown - -def toStr : FirstTokens → String - | epsilon => "epsilon" - | unknown => "unknown" - | tokens tks => toString tks - | optTokens tks => "?" ++ toString tks - -instance : ToString FirstTokens where - toString := toStr - -end FirstTokens - -structure ParserInfo where - collectTokens : List Token → List Token := id - collectKinds : SyntaxNodeKindSet → SyntaxNodeKindSet := id - firstTokens : FirstTokens := FirstTokens.unknown - deriving Inhabited - -structure Parser where - info : ParserInfo := {} - fn : ParserFn - deriving Inhabited - -abbrev TrailingParser := Parser +namespace Lean.Parser def dbgTraceStateFn (label : String) (p : ParserFn) : ParserFn := fun c s => @@ -335,9 +68,7 @@ def dbgTraceStateFn (label : String) (p : ParserFn) : ParserFn := out: {s'.stxStack.extract sz s'.stxStack.size}" s' -def dbgTraceState (label : String) (p : Parser) : Parser where - fn := dbgTraceStateFn label p.fn - info := p.info +def dbgTraceState (label : String) : Parser → Parser := withFn (dbgTraceStateFn label) @[noinline]def epsilonInfo : ParserInfo := { firstTokens := FirstTokens.epsilon } @@ -442,26 +173,15 @@ def setLhsPrec (prec : Nat) : Parser := { fn := setLhsPrecFn prec } -private def addQuotDepthFn (i : Int) (p : ParserFn) : ParserFn := fun c s => - p { c with quotDepth := c.quotDepth + i |>.toNat } s +private def addQuotDepth (i : Int) (p : Parser) : Parser := + adaptCacheableContext (fun c => { c with quotDepth := c.quotDepth + i |>.toNat }) p -def incQuotDepth (p : Parser) : Parser := { - info := p.info - fn := addQuotDepthFn 1 p.fn -} +def incQuotDepth (p : Parser) : Parser := addQuotDepth 1 p -def decQuotDepth (p : Parser) : Parser := { - info := p.info - fn := addQuotDepthFn (-1) p.fn -} +def decQuotDepth (p : Parser) : Parser := addQuotDepth (-1) p -def suppressInsideQuotFn (p : ParserFn) : ParserFn := fun c s => - p { c with suppressInsideQuot := true } s - -def suppressInsideQuot (p : Parser) : Parser := { - info := p.info - fn := suppressInsideQuotFn p.fn -} +def suppressInsideQuot (p : Parser) : Parser := + adaptCacheableContext ({ · with suppressInsideQuot := true }) p def leadingNode (n : SyntaxNodeKind) (prec : Nat) (p : Parser) : Parser := checkPrec prec >> node n p >> setLhsPrec prec @@ -551,10 +271,7 @@ def atomicFn (p : ParserFn) : ParserFn := fun c s => | ⟨stack, lhsPrec, _, cache, some msg⟩ => ⟨stack, lhsPrec, iniPos, cache, some msg⟩ | other => other -def atomic (p : Parser) : Parser := { - info := p.info - fn := atomicFn p.fn -} +def atomic : Parser → Parser := withFn atomicFn def optionalFn (p : ParserFn) : ParserFn := fun c s => let iniSz := s.stackSize @@ -580,10 +297,7 @@ def lookaheadFn (p : ParserFn) : ParserFn := fun c s => let s := p c s if s.hasError then s else s.restore iniSz iniPos -def lookahead (p : Parser) : Parser := { - info := p.info - fn := lookaheadFn p.fn -} +def lookahead : Parser → Parser := withFn lookaheadFn def notFollowedByFn (p : ParserFn) (msg : String) : ParserFn := fun c s => let iniSz := s.stackSize @@ -626,10 +340,7 @@ def many1Fn (p : ParserFn) : ParserFn := fun c s => let s := andthenFn p (manyAux p) c s s.mkNode nullKind iniSz -def many1NoAntiquot (p : Parser) : Parser := { - info := p.info - fn := many1Fn p.fn -} +def many1NoAntiquot : Parser → Parser := withFn many1Fn private partial def sepByFnAux (p : ParserFn) (sep : ParserFn) (allowTrailingSep : Bool) (iniSz : Nat) (pOpt : Bool) : ParserFn := let rec parse (pOpt : Bool) (c s) := Id.run do @@ -727,6 +438,7 @@ def takeWhileFn (p : Char → Bool) : ParserFn := def takeWhile1Fn (p : Char → Bool) (errorMsg : String) : ParserFn := andthenFn (satisfyFn p errorMsg) (takeWhileFn p) +variable (pushMissingOnError : Bool) in partial def finishCommentBlock (nesting : Nat) : ParserFn := fun c s => let input := c.input let i := s.pos @@ -751,7 +463,7 @@ partial def finishCommentBlock (nesting : Nat) : ParserFn := fun c s => else finishCommentBlock nesting c (s.setPos i) else finishCommentBlock nesting c (s.setPos i) where - eoi s := s.mkUnexpectedError "unterminated comment" + eoi s := s.mkUnexpectedError (pushMissing := pushMissingOnError) "unterminated comment" /-- Consume whitespace and comments -/ partial def whitespace : ParserFn := fun c s => @@ -761,7 +473,7 @@ partial def whitespace : ParserFn := fun c s => else let curr := input.get' i h if curr == '\t' then - s.mkUnexpectedError "tabs are not allowed; please configure your editor to expand them" + s.mkUnexpectedError (pushMissing := false) "tabs are not allowed; please configure your editor to expand them" else if curr.isWhitespace then whitespace c (s.next' input i h) else if curr == '-' then let i := input.next' i h @@ -775,7 +487,7 @@ partial def whitespace : ParserFn := fun c s => let i := input.next i let curr := input.get i if curr == '-' || curr == '!' then s -- "/--" and "/-!" doc comment are actual tokens - else andthenFn (finishCommentBlock 1) whitespace c (s.next input i) + else andthenFn (finishCommentBlock (pushMissingOnError := false) 1) whitespace c (s.next input i) else s else s @@ -843,8 +555,10 @@ def isQuotableCharDefault (c : Char) : Bool := def quotedCharFn : ParserFn := quotedCharCoreFn isQuotableCharDefault -/-- Push `(Syntax.node tk )` into syntax stack -/ -def mkNodeToken (n : SyntaxNodeKind) (startPos : String.Pos) : ParserFn := fun c s => +/-- Push `(Syntax.node tk )` onto syntax stack if parse was successful. -/ +def mkNodeToken (n : SyntaxNodeKind) (startPos : String.Pos) : ParserFn := fun c s => Id.run do + if s.hasError then + return s let input := c.input let stopPos := s.pos let leading := mkEmptySubstringAt input startPos @@ -1080,14 +794,14 @@ private def tokenFnAux : ParserFn := fun c s => let (_, tk) := c.tokens.matchPrefix input i identFnAux i tk .anonymous c s -private def updateCache (startPos : String.Pos) (s : ParserState) : ParserState := +private def updateTokenCache (startPos : String.Pos) (s : ParserState) : ParserState := -- do not cache token parsing errors, which are rare and usually fatal and thus not worth an extra field in `TokenCache` match s with - | ⟨stack, lhsPrec, pos, _, none⟩ => + | ⟨stack, lhsPrec, pos, ⟨_, catCache⟩, none⟩ => if stack.size == 0 then s else let tk := stack.back - ⟨stack, lhsPrec, pos, { tokenCache := { startPos := startPos, stopPos := pos, token := tk } }, none⟩ + ⟨stack, lhsPrec, pos, ⟨{ startPos := startPos, stopPos := pos, token := tk }, catCache⟩, none⟩ | other => other def tokenFn (expected : List String := []) : ParserFn := fun c s => @@ -1101,7 +815,7 @@ def tokenFn (expected : List String := []) : ParserFn := fun c s => s.setPos tkc.stopPos else let s := tokenFnAux c s - updateCache i s + updateTokenCache i s def peekTokenAux (c : ParserContext) (s : ParserState) : ParserState × Except ParserState Syntax := let iniSz := s.stackSize @@ -1237,8 +951,8 @@ def checkLinebreakBefore (errorMsg : String := "line break") : Parser := { fn := checkLinebreakBeforeFn errorMsg } -private def pickNonNone (stack : Array Syntax) : Syntax := - match stack.findRev? fun stx => !stx.isNone with +private def pickNonNone (stack : SyntaxStack) : Syntax := + match stack.toSubarray.findRev? fun stx => !stx.isNone with | none => Syntax.missing | some stx => stx @@ -1358,7 +1072,7 @@ def identEq (id : Name) : Parser := { namespace ParserState -def keepTop (s : Array Syntax) (startStackSize : Nat) : Array Syntax := +def keepTop (s : SyntaxStack) (startStackSize : Nat) : SyntaxStack := let node := s.back s.shrink startStackSize |>.push node @@ -1373,8 +1087,7 @@ def keepPrevError (s : ParserState) (oldStackSize : Nat) (oldStopPos : String.Po def mergeErrors (s : ParserState) (oldStackSize : Nat) (oldError : Error) : ParserState := match s with | ⟨stack, lhsPrec, pos, cache, some err⟩ => - if oldError == err then s - else ⟨stack.shrink oldStackSize, lhsPrec, pos, cache, some (oldError.merge err)⟩ + ⟨stack.shrink oldStackSize, lhsPrec, pos, cache, if oldError == err then some err else some (oldError.merge err)⟩ | other => other def keepLatest (s : ParserState) (startStackSize : Nat) : ParserState := @@ -1517,34 +1230,21 @@ def checkLineEqFn (errorMsg : String) : ParserFn := fun c s => def checkLineEq (errorMsg : String := "checkLineEq") : Parser := { fn := checkLineEqFn errorMsg } -def withPosition (p : Parser) : Parser := { - info := p.info - fn := fun c s => - p.fn { c with savedPos? := s.pos } s -} +def withPosition : Parser → Parser := withFn fun f c s => + adaptCacheableContextFn ({ · with savedPos? := s.pos }) f c s -def withPositionAfterLinebreak (p : Parser) : Parser := { - info := p.info - fn := fun c s => - let prev := s.stxStack.back - let c := if checkTailLinebreak prev then { c with savedPos? := s.pos } else c - p.fn c s -} +def withPositionAfterLinebreak : Parser → Parser := withFn fun f c s => + let prev := s.stxStack.back + adaptCacheableContextFn (fun c => if checkTailLinebreak prev then { c with savedPos? := s.pos } else c) f c s -def withoutPosition (p : Parser) : Parser := { - info := p.info - fn := fun c s => p.fn { c with savedPos? := none } s -} +def withoutPosition (p : Parser) : Parser := + adaptCacheableContext ({ · with savedPos? := none }) p -def withForbidden (tk : Token) (p : Parser) : Parser := { - info := p.info - fn := fun c s => p.fn { c with forbiddenTk? := tk } s -} +def withForbidden (tk : Token) (p : Parser) : Parser := + adaptCacheableContext ({ · with forbiddenTk? := tk }) p -def withoutForbidden (p : Parser) : Parser := { - info := p.info - fn := fun c s => p.fn { c with forbiddenTk? := none } s -} +def withoutForbidden (p : Parser) : Parser := + adaptCacheableContext ({ · with forbiddenTk? := none }) p def eoiFn : ParserFn := fun c s => let i := s.pos @@ -1686,9 +1386,8 @@ builtin_initialize categoryParserFnExtension : EnvExtension CategoryParserFn ← def categoryParserFn (catName : Name) : ParserFn := fun ctx s => categoryParserFnExtension.getState ctx.env catName ctx s -def categoryParser (catName : Name) (prec : Nat) : Parser := { - fn := fun c s => categoryParserFn catName { c with prec := prec } s -} +def categoryParser (catName : Name) (prec : Nat) : Parser where + fn := adaptCacheableContextFn ({ · with prec }) (withCacheFn catName (categoryParserFn catName)) -- Define `termParser` here because we need it for antiquotations def termParser (prec : Nat := 0) : Parser := @@ -1719,9 +1418,7 @@ def setExpectedFn (expected : List String) (p : ParserFn) : ParserFn := fun c s | s'@{ errorMsg := some msg, .. } => { s' with errorMsg := some { msg with expected } } | s' => s' -def setExpected (expected : List String) (p : Parser) : Parser := { - fn := setExpectedFn expected p.fn, info := p.info -} +def setExpected (expected : List String) : Parser → Parser := withFn (setExpectedFn expected) def pushNone : Parser := { fn := fun _ s => s.pushSyntax mkNullNode @@ -1741,15 +1438,13 @@ def tokenAntiquotFn : ParserFn := fun c s => Id.run do return s.restore iniSz iniPos s.mkNode (`token_antiquot) (iniSz - 1) -def tokenWithAntiquot (p : Parser) : Parser where - fn c s := - let s := p.fn c s - -- fast check that is false in most cases - if c.input.get s.pos == '%' then - tokenAntiquotFn c s - else - s - info := p.info +def tokenWithAntiquot : Parser → Parser := withFn fun f c s => + let s := f c s + -- fast check that is false in most cases + if c.input.get s.pos == '%' then + tokenAntiquotFn c s + else + s def symbol (sym : String) : Parser := tokenWithAntiquot (symbolNoAntiquot sym) @@ -1849,23 +1544,6 @@ def sepBy (p : Parser) (sep : String) (psep : Parser := symbol sep) (allowTraili def sepBy1 (p : Parser) (sep : String) (psep : Parser := symbol sep) (allowTrailingSep : Bool := false) : Parser := sepBy1NoAntiquot (sepByElemParser p sep) psep allowTrailingSep -def categoryParserOfStackFn (offset : Nat) : ParserFn := fun ctx s => - let stack := s.stxStack - if h : stack.size < offset + 1 then - s.mkUnexpectedError ("failed to determine parser category using syntax stack, stack is too small") - else - have : stack.size - (offset + 1) < stack.size := by - apply Nat.sub_lt <;> simp_all_arith - apply Nat.le_trans (m := offset + 1) - apply Nat.le_add_left; assumption - match stack[stack.size - (offset + 1)] with - | .ident _ _ catName _ => categoryParserFn catName ctx s - | _ => s.mkUnexpectedError ("failed to determine parser category using syntax stack, the specified element on the stack is not an identifier") - -def categoryParserOfStack (offset : Nat) (prec : Nat := 0) : Parser := { - fn := fun c s => categoryParserOfStackFn offset { c with prec := prec } s -} - private def mkResult (s : ParserState) (iniSz : Nat) : ParserState := if s.stackSize == iniSz + 1 then s else s.mkNode nullKind iniSz -- throw error instead? diff --git a/src/Lean/Parser/Extension.lean b/src/Lean/Parser/Extension.lean index a39aacb152a9..5165eb64c593 100644 --- a/src/Lean/Parser/Extension.lean +++ b/src/Lean/Parser/Extension.lean @@ -280,7 +280,7 @@ partial def compileParserDescr (categories : ParserCategories) (d : ParserDescr) | ParserDescr.unary n d => return (← getUnaryAlias parserAliasesRef n) (← visit d) | ParserDescr.binary n d₁ d₂ => return (← getBinaryAlias parserAliasesRef n) (← visit d₁) (← visit d₂) | ParserDescr.node k prec d => return leadingNode k prec (← visit d) - | ParserDescr.nodeWithAntiquot n k d => return nodeWithAntiquot n k (← visit d) (anonymous := true) + | ParserDescr.nodeWithAntiquot n k d => return withCache k (nodeWithAntiquot n k (← visit d) (anonymous := true)) | ParserDescr.sepBy p sep psep trail => return sepBy (← visit p) sep (← visit psep) trail | ParserDescr.sepBy1 p sep psep trail => return sepBy1 (← visit p) sep (← visit psep) trail | ParserDescr.trailingNode k prec lhsPrec d => return trailingNode k prec lhsPrec (← visit d) @@ -364,8 +364,7 @@ unsafe def evalParserConstUnsafe (declName : Name) : ParserFn := fun ctx s => un match (← (mkParserOfConstant categories declName { env := ctx.env, opts := ctx.options }).toBaseIO) with | .ok (_, p) => -- We should manually register `p`'s tokens before invoking it as it might not be part of any syntax category (yet) - let ctx := { ctx with tokens := p.info.collectTokens [] |>.foldl (fun tks tk => tks.insert tk tk) ctx.tokens } - return p.fn ctx s + return adaptUncacheableContextFn (fun ctx => { ctx with tokens := p.info.collectTokens [] |>.foldl (fun tks tk => tks.insert tk tk) ctx.tokens }) p.fn ctx s | .error e => return s.mkUnexpectedError e.toString @[implemented_by evalParserConstUnsafe] @@ -378,12 +377,11 @@ register_builtin_option internal.parseQuotWithCurrentStage : Bool := { } /-- Run `declName` if possible and inside a quotation, or else `p`. The `ParserInfo` will always be taken from `p`. -/ -def evalInsideQuot (declName : Name) (p : Parser) : Parser := { p with - fn := fun c s => - if c.quotDepth > 0 && !c.suppressInsideQuot && internal.parseQuotWithCurrentStage.get c.options && c.env.contains declName then - evalParserConst declName c s - else - p.fn c s } +def evalInsideQuot (declName : Name) : Parser → Parser := withFn fun f c s => + if c.quotDepth > 0 && !c.suppressInsideQuot && internal.parseQuotWithCurrentStage.get c.options && c.env.contains declName then + evalParserConst declName c s + else + f c s def addBuiltinParser (catName : Name) (declName : Name) (leading : Bool) (p : Parser) (prio : Nat) : IO Unit := do let p := evalInsideQuot declName p @@ -445,28 +443,20 @@ def mkInputContext (input : String) (fileName : String) : InputContext := { fileMap := input.toFileMap } -def mkParserContext (ictx : InputContext) (pmctx : ParserModuleContext) : ParserContext := { - prec := 0, - toInputContext := ictx, - toParserModuleContext := pmctx, - tokens := getTokenTable pmctx.env -} - def mkParserState (input : String) : ParserState := { cache := initCacheForInput input } /-- convenience function for testing -/ def runParserCategory (env : Environment) (catName : Name) (input : String) (fileName := "") : Except String Syntax := - let c := mkParserContext (mkInputContext input fileName) { env := env, options := {} } - let s := mkParserState input - let s := whitespace c s - let s := categoryParserFnImpl catName c s + let p := andthenFn whitespace (categoryParserFnImpl catName) + let ictx := mkInputContext input fileName + let s := p.run ictx { env, options := {} } (getTokenTable env) (mkParserState input) if s.hasError then - Except.error (s.toErrorMsg c) + Except.error (s.toErrorMsg ictx) else if input.atEnd s.pos then Except.ok s.stxStack.back else - Except.error ((s.mkError "end of input").toErrorMsg c) + Except.error ((s.mkError "end of input").toErrorMsg ictx) def declareBuiltinParser (addFnName : Name) (catName : Name) (declName : Name) (prio : Nat) : CoreM Unit := let val := mkAppN (mkConst addFnName) #[toExpr catName, toExpr declName, mkConst declName, mkRawNatLit prio] @@ -578,7 +568,7 @@ builtin_initialize registerBuiltinDynamicParserAttribute `command_parser `comman @[inline] def commandParser (rbp : Nat := 0) : Parser := categoryParser `command rbp -private def withNamespaces (ids : Array Name) (p : ParserFn) (addOpenSimple : Bool) : ParserFn := fun c s => +private def withNamespaces (ids : Array Name) (addOpenSimple : Bool) : ParserFn → ParserFn := adaptUncacheableContextFn fun c => let c := ids.foldl (init := c) fun c id => let nss := ResolveName.resolveNamespace c.env c.currNamespace c.openDecls id let (env, openDecls) := nss.foldl (init := (c.env, c.openDecls)) fun (env, openDecls) ns => @@ -587,7 +577,7 @@ private def withNamespaces (ids : Array Name) (p : ParserFn) (addOpenSimple : Bo (env, openDecls) { c with env, openDecls } let tokens := parserExtension.getState c.env |>.tokens - p { c with tokens } s + { c with tokens } def withOpenDeclFnCore (openDeclStx : Syntax) (p : ParserFn) : ParserFn := fun c s => if openDeclStx.getKind == `Lean.Parser.Command.openSimple then @@ -615,10 +605,7 @@ def withOpenFn (p : ParserFn) : ParserFn := fun c s => p c s -@[inline] def withOpen (p : Parser) : Parser := { - info := p.info - fn := withOpenFn p.fn -} +@[inline] def withOpen : Parser → Parser := withFn withOpenFn /-- If the parsing stack is of the form `#[.., openDecl]`, we process the open declaration, and execute `p` -/ def withOpenDeclFn (p : ParserFn) : ParserFn := fun c s => @@ -628,10 +615,7 @@ def withOpenDeclFn (p : ParserFn) : ParserFn := fun c s => else p c s -@[inline] def withOpenDecl (p : Parser) : Parser := { - info := p.info - fn := withOpenDeclFn p.fn -} +@[inline] def withOpenDecl : Parser → Parser := withFn withOpenDeclFn inductive ParserName | category (cat : Name) @@ -684,12 +668,12 @@ def parserOfStackFn (offset : Nat) : ParserFn := fun ctx s => Id.run do categoryParserFn cat ctx s | [.parser parserName _] => let iniSz := s.stackSize - let mut ctx' := ctx - if !internal.parseQuotWithCurrentStage.get ctx'.options then - -- static quotations such as `(e) do not use the interpreter unless the above option is set, - -- so for consistency neither should dynamic quotations using this function - ctx' := { ctx' with options := ctx'.options.setBool `interpreter.prefer_native true } - let s := evalParserConst parserName ctx' s + let s := adaptUncacheableContextFn (fun ctx => + if !internal.parseQuotWithCurrentStage.get ctx.options then + -- static quotations such as `(e) do not use the interpreter unless the above option is set, + -- so for consistency neither should dynamic quotations using this function + { ctx with options := ctx.options.setBool `interpreter.prefer_native true } + else ctx) (evalParserConst parserName) ctx s if !s.hasError && s.stackSize != iniSz + 1 then s.mkUnexpectedError "expected parser to return exactly one syntax object" else @@ -697,8 +681,8 @@ def parserOfStackFn (offset : Nat) : ParserFn := fun ctx s => Id.run do | _::_::_ => s.mkUnexpectedError s!"ambiguous parser name {parserName}" | [] => s.mkUnexpectedError s!"unknown parser {parserName}" -def parserOfStack (offset : Nat) (prec : Nat := 0) : Parser := - { fn := fun c s => parserOfStackFn offset { c with prec := prec } s } +def parserOfStack (offset : Nat) (prec : Nat := 0) : Parser where + fn := adaptCacheableContextFn ({ · with prec }) (parserOfStackFn offset) end Parser end Lean diff --git a/src/Lean/Parser/Extra.lean b/src/Lean/Parser/Extra.lean index 80b62cbd1617..53ee5b427b83 100644 --- a/src/Lean/Parser/Extra.lean +++ b/src/Lean/Parser/Extra.lean @@ -16,6 +16,10 @@ namespace Parser attribute [run_builtin_parser_attribute_hooks] leadingNode termParser commandParser mkAntiquot nodeWithAntiquot sepBy sepBy1 unicodeSymbol nonReservedSymbol + withCache withPosition withPositionAfterLinebreak withoutPosition withForbidden withoutForbidden setExpected + incQuotDepth decQuotDepth suppressInsideQuot evalInsideQuot + withOpen withOpenDecl + dbgTraceState @[run_builtin_parser_attribute_hooks] def optional (p : Parser) : Parser := optionalNoAntiquot (withAntiquotSpliceAndSuffix `optional p (symbol "?")) diff --git a/src/Lean/Parser/Module.lean b/src/Lean/Parser/Module.lean index c02219d85b71..b7afc0e118d6 100644 --- a/src/Lean/Parser/Module.lean +++ b/src/Lean/Parser/Module.lean @@ -20,11 +20,10 @@ def header := leading_parser optional («prelude» >> ppLine) >> many («imp @[run_builtin_parser_attribute_hooks] def module := leading_parser header >> many (commandParser >> ppLine >> ppLine) -def updateTokens (c : ParserContext) : ParserContext := - { c with - tokens := match addParserTokens c.tokens header.info with - | Except.ok tables => tables - | Except.error _ => unreachable! } +def updateTokens (tokens : TokenTable) : TokenTable := + match addParserTokens tokens header.info with + | Except.ok tables => tables + | Except.error _ => unreachable! end Module @@ -33,21 +32,19 @@ structure ModuleParserState where recovering : Bool := false deriving Inhabited -private def mkErrorMessage (c : ParserContext) (pos : String.Pos) (errorMsg : String) : Message := +private def mkErrorMessage (c : InputContext) (pos : String.Pos) (errorMsg : String) : Message := let pos := c.fileMap.toPosition pos { fileName := c.fileName, pos := pos, data := errorMsg } def parseHeader (inputCtx : InputContext) : IO (Syntax × ModuleParserState × MessageLog) := do let dummyEnv ← mkEmptyEnvironment - let ctx := mkParserContext inputCtx { env := dummyEnv, options := {} } - let ctx := Module.updateTokens ctx - let s := mkParserState ctx.input - let s := whitespace ctx s - let s := Module.header.fn ctx s + let p := andthenFn whitespace Module.header.fn + let tokens := Module.updateTokens (getTokenTable dummyEnv) + let s := p.run inputCtx { env := dummyEnv, options := {} } tokens (mkParserState inputCtx.input) let stx := s.stxStack.back match s.errorMsg with | some errorMsg => - let msg := mkErrorMessage ctx s.pos (toString errorMsg) + let msg := mkErrorMessage inputCtx s.pos (toString errorMsg) pure (stx, { pos := s.pos, recovering := true }, { : MessageLog }.add msg) | none => pure (stx, { pos := s.pos }, {}) @@ -62,9 +59,9 @@ def isEOI (s : Syntax) : Bool := def isTerminalCommand (s : Syntax) : Bool := s.isOfKind ``Command.exit || s.isOfKind ``Command.import -private def consumeInput (c : ParserContext) (pos : String.Pos) : String.Pos := - let s : ParserState := { cache := initCacheForInput c.input, pos := pos } - let s := tokenFn [] c s +private def consumeInput (inputCtx : InputContext) (pmctx : ParserModuleContext) (pos : String.Pos) : String.Pos := + let s : ParserState := { cache := initCacheForInput inputCtx.input, pos := pos } + let s := tokenFn [] |>.run inputCtx pmctx (getTokenTable pmctx.env) s match s.errorMsg with | some _ => pos + ' ' | none => s.pos @@ -82,10 +79,8 @@ partial def parseCommand (inputCtx : InputContext) (pmctx : ParserModuleContext) stx := mkEOI pos break let pos' := pos - let c := mkParserContext inputCtx pmctx - let s := { cache := initCacheForInput c.input, pos := pos : ParserState } - let s := whitespace c s - let s := topLevelCommandParserFn c s + let p := andthenFn whitespace topLevelCommandParserFn + let s := p.run inputCtx pmctx (getTokenTable pmctx.env) { cache := initCacheForInput inputCtx.input, pos } pos := s.pos if recovering && !s.stxStack.isEmpty && s.stxStack.back.isAntiquot then -- top-level antiquotation during recovery is most likely remnant from unfinished quotation, ignore @@ -98,13 +93,13 @@ partial def parseCommand (inputCtx : InputContext) (pmctx : ParserModuleContext) | some errorMsg => -- advance at least one token to prevent infinite loops if pos == pos' then - pos := consumeInput c pos + pos := consumeInput inputCtx pmctx pos /- We ignore commands where `getPos?` is none. This happens only on commands that have a prefix comprised of optional elements. For example, unification hints start with `optional («scoped» <|> «local»)`. We claim a syntactically incorrect command containing no token or identifier is irrelevant for intellisense and should be ignored. -/ let ignore := s.stxStack.isEmpty || s.stxStack.back.getPos?.isNone unless recovering && ignore do - messages := messages.add <| mkErrorMessage c s.pos (toString errorMsg) + messages := messages.add <| mkErrorMessage inputCtx s.pos (toString errorMsg) recovering := true if ignore then continue diff --git a/src/Lean/Parser/Term.lean b/src/Lean/Parser/Term.lean index c5053baccfa5..f972835ebaab 100644 --- a/src/Lean/Parser/Term.lean +++ b/src/Lean/Parser/Term.lean @@ -11,7 +11,7 @@ namespace Parser namespace Command def commentBody : Parser := -{ fn := rawFn (finishCommentBlock 1) (trailingWs := true) } +{ fn := rawFn (finishCommentBlock (pushMissingOnError := true) 1) (trailingWs := true) } @[combinator_parenthesizer commentBody] def commentBody.parenthesizer := PrettyPrinter.Parenthesizer.visitToken @@ -660,8 +660,8 @@ def macroLastArg := macroDollarArg <|> macroArg @[builtin_term_parser] def stateRefT := leading_parser "StateRefT" >> macroArg >> macroLastArg -@[builtin_term_parser] def dynamicQuot := leading_parser - "`(" >> ident >> "|" >> withoutPosition (incQuotDepth (parserOfStack 1)) >> ")" +@[builtin_term_parser] def dynamicQuot := withoutPosition <| leading_parser + "`(" >> ident >> "|" >> incQuotDepth (parserOfStack 1) >> ")" @[builtin_term_parser] def dotIdent := leading_parser "." >> checkNoWsBefore >> rawIdent diff --git a/src/Lean/Parser/Types.lean b/src/Lean/Parser/Types.lean new file mode 100644 index 000000000000..f9e7adc1c32e --- /dev/null +++ b/src/Lean/Parser/Types.lean @@ -0,0 +1,395 @@ +/- +Copyright (c) 2019 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Sebastian Ullrich +-/ +import Lean.Data.Trie +import Lean.Syntax +import Lean.Message + +namespace Lean.Parser + +abbrev mkAtom (info : SourceInfo) (val : String) : Syntax := + Syntax.atom info val + +abbrev mkIdent (info : SourceInfo) (rawVal : Substring) (val : Name) : Syntax := + Syntax.ident info rawVal val [] + +/-- Return character after position `pos` -/ +def getNext (input : String) (pos : String.Pos) : Char := + input.get (input.next pos) + +/-- Maximal (and function application) precedence. + In the standard lean language, no parser has precedence higher than `maxPrec`. + + Note that nothing prevents users from using a higher precedence, but we strongly + discourage them from doing it. -/ +def maxPrec : Nat := eval_prec max +def argPrec : Nat := eval_prec arg +def leadPrec : Nat := eval_prec lead +def minPrec : Nat := eval_prec min + +abbrev Token := String + +abbrev TokenTable := Trie Token + +abbrev SyntaxNodeKindSet := PersistentHashMap SyntaxNodeKind Unit + +def SyntaxNodeKindSet.insert (s : SyntaxNodeKindSet) (k : SyntaxNodeKind) : SyntaxNodeKindSet := + PersistentHashMap.insert s k () + +/-- + Input string and related data. Recall that the `FileMap` is a helper structure for mapping + `String.Pos` in the input string to line/column information. -/ +structure InputContext where + input : String + fileName : String + fileMap : FileMap + deriving Inhabited + +/-- Input context derived from elaboration of previous commands. -/ +structure ParserModuleContext where + env : Environment + options : Options + -- for name lookup + currNamespace : Name := .anonymous + openDecls : List OpenDecl := [] + +/-- Parser context parts that can be updated without invalidating the parser cache. -/ +structure CacheableParserContext where + prec : Nat + -- used for bootstrapping only + quotDepth : Nat := 0 + suppressInsideQuot : Bool := false + savedPos? : Option String.Pos := none + forbiddenTk? : Option Token := none + deriving BEq + +/-- Parser context updateable in `adaptUncacheableContextFn`. -/ +structure ParserContextCore extends InputContext, ParserModuleContext, CacheableParserContext where + tokens : TokenTable + +/-- Opaque parser context updateable using `adaptCacheableContextFn` and `adaptUncacheableContextFn`. -/ +structure ParserContext extends ParserContextCore where private mk :: + +structure Error where + unexpected : String := "" + expected : List String := [] + deriving Inhabited, BEq + +namespace Error + +private def expectedToString : List String → String + | [] => "" + | [e] => e + | [e1, e2] => e1 ++ " or " ++ e2 + | e::es => e ++ ", " ++ expectedToString es + +protected def toString (e : Error) : String := + let unexpected := if e.unexpected == "" then [] else [e.unexpected] + let expected := if e.expected == [] then [] else + let expected := e.expected.toArray.qsort (fun e e' => e < e') + let expected := expected.toList.eraseReps + ["expected " ++ expectedToString expected] + "; ".intercalate $ unexpected ++ expected + +instance : ToString Error where + toString := Error.toString + +def merge (e₁ e₂ : Error) : Error := + match e₂ with + | { unexpected := u, .. } => { unexpected := if u == "" then e₁.unexpected else u, expected := e₁.expected ++ e₂.expected } + +end Error + +structure TokenCacheEntry where + startPos : String.Pos := 0 + stopPos : String.Pos := 0 + token : Syntax := Syntax.missing + +structure ParserCacheKey extends CacheableParserContext where + parserName : Name + pos : String.Pos + deriving BEq + +instance : Hashable ParserCacheKey where + -- sufficient to avoid most collisions, relatively cheap to compute + hash k := hash (k.pos, k.parserName) + +structure ParserCacheEntry where + stx : Syntax + lhsPrec : Nat + newPos : String.Pos + errorMsg : Option Error + +structure ParserCache where + tokenCache : TokenCacheEntry + parserCache : HashMap ParserCacheKey ParserCacheEntry + +def initCacheForInput (input : String) : ParserCache where + tokenCache := { startPos := input.endPos + ' ' /- make sure it is not a valid position -/ } + parserCache := {} + +/-- A syntax array with an inaccessible prefix, used for sound caching. -/ +structure SyntaxStack where + private raw : Array Syntax + private drop : Nat + +namespace SyntaxStack + +def toSubarray (stack : SyntaxStack) : Subarray Syntax := + stack.raw.toSubarray stack.drop + +def empty : SyntaxStack where + raw := #[] + drop := 0 + +def size (stack : SyntaxStack) : Nat := + stack.raw.size - stack.drop + +def isEmpty (stack : SyntaxStack) : Bool := + stack.size == 0 + +def shrink (stack : SyntaxStack) (n : Nat) : SyntaxStack := + { stack with raw := stack.raw.shrink (stack.drop + n) } + +def push (stack : SyntaxStack) (a : Syntax) : SyntaxStack := + { stack with raw := stack.raw.push a } + +def pop (stack : SyntaxStack) : SyntaxStack := + if stack.size > 0 then + { stack with raw := stack.raw.pop } + else stack + +def back (stack : SyntaxStack) : Syntax := + if stack.size > 0 then + stack.raw.back + else + panic! "SyntaxStack.back: element is inaccessible" + +def get! (stack : SyntaxStack) (i : Nat) : Syntax := + if i < stack.size then + stack.raw.get! (stack.drop + i) + else + panic! "SyntaxStack.get!: element is inaccessible" + +def extract (stack : SyntaxStack) (start stop : Nat) : Array Syntax := + stack.raw.extract (stack.drop + start) (stack.drop + stop) + +instance : HAppend SyntaxStack (Array Syntax) SyntaxStack where + hAppend stack stxs := { stack with raw := stack.raw ++ stxs } + +end SyntaxStack + +structure ParserState where + stxStack : SyntaxStack := .empty + /-- + Set to the precedence of the preceding (not surrounding) parser by `runLongestMatchParser` + for the use of `checkLhsPrec` in trailing parsers. + Note that with chaining, the preceding parser can be another trailing parser: + in `1 * 2 + 3`, the preceding parser is '*' when '+' is executed. -/ + lhsPrec : Nat := 0 + pos : String.Pos := 0 + cache : ParserCache + errorMsg : Option Error := none + +namespace ParserState + +def hasError (s : ParserState) : Bool := + s.errorMsg != none + +def stackSize (s : ParserState) : Nat := + s.stxStack.size + +def restore (s : ParserState) (iniStackSz : Nat) (iniPos : String.Pos) : ParserState := + { s with stxStack := s.stxStack.shrink iniStackSz, errorMsg := none, pos := iniPos } + +def setPos (s : ParserState) (pos : String.Pos) : ParserState := + { s with pos := pos } + +def setCache (s : ParserState) (cache : ParserCache) : ParserState := + { s with cache := cache } + +def pushSyntax (s : ParserState) (n : Syntax) : ParserState := + { s with stxStack := s.stxStack.push n } + +def popSyntax (s : ParserState) : ParserState := + { s with stxStack := s.stxStack.pop } + +def shrinkStack (s : ParserState) (iniStackSz : Nat) : ParserState := + { s with stxStack := s.stxStack.shrink iniStackSz } + +def next (s : ParserState) (input : String) (pos : String.Pos) : ParserState := + { s with pos := input.next pos } + +def next' (s : ParserState) (input : String) (pos : String.Pos) (h : ¬ input.atEnd pos) : ParserState := + { s with pos := input.next' pos h } + +def toErrorMsg (ctx : InputContext) (s : ParserState) : String := + match s.errorMsg with + | none => "" + | some msg => + let pos := ctx.fileMap.toPosition s.pos + mkErrorStringWithPos ctx.fileName pos (toString msg) + +def mkNode (s : ParserState) (k : SyntaxNodeKind) (iniStackSz : Nat) : ParserState := + match s with + | ⟨stack, lhsPrec, pos, cache, err⟩ => + if err != none && stack.size == iniStackSz then + -- If there is an error but there are no new nodes on the stack, use `missing` instead. + -- Thus we ensure the property that an syntax tree contains (at least) one `missing` node + -- if (and only if) there was a parse error. + -- We should not create an actual node of kind `k` in this case because it would mean we + -- choose an "arbitrary" node (in practice the last one) in an alternative of the form + -- `node k1 p1 <|> ... <|> node kn pn` when all parsers fail. With the code below we + -- instead return a less misleading single `missing` node without randomly selecting any `ki`. + let stack := stack.push Syntax.missing + ⟨stack, lhsPrec, pos, cache, err⟩ + else + let newNode := Syntax.node SourceInfo.none k (stack.extract iniStackSz stack.size) + let stack := stack.shrink iniStackSz + let stack := stack.push newNode + ⟨stack, lhsPrec, pos, cache, err⟩ + +def mkTrailingNode (s : ParserState) (k : SyntaxNodeKind) (iniStackSz : Nat) : ParserState := + match s with + | ⟨stack, lhsPrec, pos, cache, err⟩ => + let newNode := Syntax.node SourceInfo.none k (stack.extract (iniStackSz - 1) stack.size) + let stack := stack.shrink (iniStackSz - 1) + let stack := stack.push newNode + ⟨stack, lhsPrec, pos, cache, err⟩ + +def setError (s : ParserState) (msg : String) : ParserState := + match s with + | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨stack, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ + +def mkError (s : ParserState) (msg : String) : ParserState := + match s with + | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ + +def mkUnexpectedError (s : ParserState) (msg : String) (expected : List String := []) (pushMissing := true) : ParserState := + match s with + | ⟨stack, lhsPrec, pos, cache, _⟩ => ⟨if pushMissing then stack.push .missing else stack, lhsPrec, pos, cache, some { unexpected := msg, expected := expected }⟩ + +def mkEOIError (s : ParserState) (expected : List String := []) : ParserState := + s.mkUnexpectedError "unexpected end of input" expected + +def mkErrorAt (s : ParserState) (msg : String) (pos : String.Pos) (initStackSz? : Option Nat := none) : ParserState := + match s, initStackSz? with + | ⟨stack, lhsPrec, _, cache, _⟩, none => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ + | ⟨stack, lhsPrec, _, cache, _⟩, some sz => ⟨stack.shrink sz |>.push Syntax.missing, lhsPrec, pos, cache, some { expected := [ msg ] }⟩ + +def mkErrorsAt (s : ParserState) (ex : List String) (pos : String.Pos) (initStackSz? : Option Nat := none) : ParserState := + match s, initStackSz? with + | ⟨stack, lhsPrec, _, cache, _⟩, none => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { expected := ex }⟩ + | ⟨stack, lhsPrec, _, cache, _⟩, some sz => ⟨stack.shrink sz |>.push Syntax.missing, lhsPrec, pos, cache, some { expected := ex }⟩ + +def mkUnexpectedErrorAt (s : ParserState) (msg : String) (pos : String.Pos) : ParserState := + match s with + | ⟨stack, lhsPrec, _, cache, _⟩ => ⟨stack.push Syntax.missing, lhsPrec, pos, cache, some { unexpected := msg }⟩ + +end ParserState + +def ParserFn := ParserContext → ParserState → ParserState + +instance : Inhabited ParserFn where + default := fun _ s => s + +inductive FirstTokens where + | epsilon : FirstTokens + | unknown : FirstTokens + | tokens : List Token → FirstTokens + | optTokens : List Token → FirstTokens + deriving Inhabited + +namespace FirstTokens + +def seq : FirstTokens → FirstTokens → FirstTokens + | epsilon, tks => tks + | optTokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) + | optTokens s₁, tokens s₂ => tokens (s₁ ++ s₂) + | tks, _ => tks + +def toOptional : FirstTokens → FirstTokens + | tokens tks => optTokens tks + | tks => tks + +def merge : FirstTokens → FirstTokens → FirstTokens + | epsilon, tks => toOptional tks + | tks, epsilon => toOptional tks + | tokens s₁, tokens s₂ => tokens (s₁ ++ s₂) + | optTokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) + | tokens s₁, optTokens s₂ => optTokens (s₁ ++ s₂) + | optTokens s₁, tokens s₂ => optTokens (s₁ ++ s₂) + | _, _ => unknown + +def toStr : FirstTokens → String + | epsilon => "epsilon" + | unknown => "unknown" + | tokens tks => toString tks + | optTokens tks => "?" ++ toString tks + +instance : ToString FirstTokens where + toString := toStr + +end FirstTokens + +structure ParserInfo where + collectTokens : List Token → List Token := id + collectKinds : SyntaxNodeKindSet → SyntaxNodeKindSet := id + firstTokens : FirstTokens := FirstTokens.unknown + deriving Inhabited + +structure Parser where + info : ParserInfo := {} + fn : ParserFn + deriving Inhabited + +abbrev TrailingParser := Parser + +/-- Create a simple parser combinator that inherits the `info` of the nested parser. -/ +@[inline] +def withFn (f : ParserFn → ParserFn) (p : Parser) : Parser := { p with fn := f p.fn } + +def adaptCacheableContextFn (f : CacheableParserContext → CacheableParserContext) (p : ParserFn) : ParserFn := fun c s => + p { c with toCacheableParserContext := f c.toCacheableParserContext } s + +def adaptCacheableContext (f : CacheableParserContext → CacheableParserContext) : Parser → Parser := + withFn (adaptCacheableContextFn f) + +/-- Run `p` under the given context transformation with a fresh cache, restore outer cache afterwards. -/ +def adaptUncacheableContextFn (f : ParserContextCore → ParserContextCore) (p : ParserFn) : ParserFn := fun c s => + let parserCache := s.cache.parserCache + let s' := p ⟨f c.toParserContextCore⟩ { s with cache.parserCache := {} } + { s' with cache.parserCache := parserCache } + +/-- +Run `p` and record result in parser cache for any further invocation with this `parserName`, parser context, and parser state. +`p` cannot access syntax stack elements pushed before the invocation in order to make caching independent of parser history. +As this excludes trailing parsers from being cached, we also reset `lhsPrec`, which is not read but set by leading parsers, to 0 +in order to increase cache hits. Finally, `errorMsg` is also reset to `none` as a leading parser should not be called in the first +place if there was an error. +-/ +def withCacheFn (parserName : Name) (p : ParserFn) : ParserFn := fun c s => Id.run do + let key := ⟨c.toCacheableParserContext, parserName, s.pos⟩ + if let some r := s.cache.parserCache.find? key then + -- TODO: turn this into a proper trace once we have these in the parser + --dbg_trace "parser cache hit: {parserName}:{s.pos} -> {r.stx}" + return ⟨s.stxStack.push r.stx, r.lhsPrec, r.newPos, s.cache, r.errorMsg⟩ + let initDrop := s.stxStack.drop + let initStackSz := s.stxStack.raw.size + let s := p c { s with stxStack.drop := initStackSz, lhsPrec := 0, errorMsg := none } + let s := { s with stxStack.drop := initDrop } + if s.stxStack.raw.size != initStackSz + 1 then + panic! s!"withCacheFn: unexpected stack growth {s.stxStack.raw}" + { s with cache.parserCache := s.cache.parserCache.insert key ⟨s.stxStack.back, s.lhsPrec, s.pos, s.errorMsg⟩ } + +@[inherit_doc withCacheFn] +def withCache (parserName : Name) : Parser → Parser := withFn (withCacheFn parserName) + +def ParserFn.run (p : ParserFn) (ictx : InputContext) (pmctx : ParserModuleContext) (tokens : TokenTable) (s : ParserState) : ParserState := + p { pmctx with + prec := 0 + toInputContext := ictx + tokens + } s diff --git a/src/Lean/ParserCompiler.lean b/src/Lean/ParserCompiler.lean index 1128960b9914..ba73523eb6a5 100644 --- a/src/Lean/ParserCompiler.lean +++ b/src/Lean/ParserCompiler.lean @@ -37,9 +37,7 @@ partial def parserNodeKind? (e : Expr) : MetaM (Option Name) := do let e ← whnfCore e if e matches Expr.lam .. then lambdaLetTelescope e fun _ e => parserNodeKind? e - else if e.isAppOfArity ``nodeWithAntiquot 4 then - reduceEval? (e.getArg! 1) - else if e.isAppOfArity ``withAntiquot 2 then + else if e.isAppOfArity ``nodeWithAntiquot 4 || e.isAppOfArity ``withAntiquot 2 || e.isAppOfArity ``withCache 2 then parserNodeKind? (e.getArg! 1) else if e.isAppOfArity ``leadingNode 3 || e.isAppOfArity ``trailingNode 4 || e.isAppOfArity ``node 2 then reduceEval? (e.getArg! 0) diff --git a/src/Lean/PrettyPrinter/Formatter.lean b/src/Lean/PrettyPrinter/Formatter.lean index 87c613c1b020..a3a6a2aaa114 100644 --- a/src/Lean/PrettyPrinter/Formatter.lean +++ b/src/Lean/PrettyPrinter/Formatter.lean @@ -267,12 +267,6 @@ def categoryParser.formatter (cat : Name) : Formatter := do def categoryFormatter (cat : Name) : Formatter := fill <| indent <| categoryFormatterCore cat -@[combinator_formatter categoryParserOfStack] -def categoryParserOfStack.formatter (offset : Nat) : Formatter := do - let st ← get - let stx := st.stxTrav.parents.back.getArg (st.stxTrav.idxs.back - offset) - categoryParser.formatter stx.getId - @[combinator_formatter parserOfStack] def parserOfStack.formatter (offset : Nat) (_prec : Nat := 0) : Formatter := do let st ← get @@ -283,8 +277,6 @@ def parserOfStack.formatter (offset : Nat) (_prec : Nat := 0) : Formatter := do def error.formatter (_msg : String) : Formatter := pure () @[combinator_formatter errorAtSavedPos] def errorAtSavedPos.formatter (_msg : String) (_delta : Bool) : Formatter := pure () -@[combinator_formatter atomic] -def atomic.formatter (p : Formatter) : Formatter := p @[combinator_formatter lookahead] def lookahead.formatter (_ : Formatter) : Formatter := pure () @@ -305,6 +297,9 @@ def node.formatter (k : SyntaxNodeKind) (p : Formatter) : Formatter := do checkKind k; visitArgs p +@[combinator_formatter withFn] +def withFn.formatter (_ : ParserFn → ParserFn) (p : Formatter) : Formatter := p + @[combinator_formatter trailingNode] def trailingNode.formatter (k : SyntaxNodeKind) (_ _ : Nat) (p : Formatter) : Formatter := do checkKind k @@ -315,14 +310,14 @@ def trailingNode.formatter (k : SyntaxNodeKind) (_ _ : Nat) (p : Formatter) : Fo def parseToken (s : String) : FormatterM ParserState := -- include comment tokens, e.g. when formatting `- -0` - return (Parser.andthenFn Parser.whitespace (Parser.tokenFn [])) { + return (Parser.andthenFn Parser.whitespace (Parser.tokenFn [])).run { input := s, fileName := "", - fileMap := FileMap.ofString "", - prec := 0, + fileMap := FileMap.ofString "" + } { env := ← getEnv, - options := ← getOptions, - tokens := (← read).table } (Parser.mkParserState s) + options := ← getOptions + } ((← read).table) (Parser.mkParserState s) def pushToken (info : SourceInfo) (tk : String) : FormatterM Unit := do match info with @@ -462,20 +457,7 @@ def sepByNoAntiquot.formatter (p pSep : Formatter) : Formatter := do @[combinator_formatter sepBy1NoAntiquot] def sepBy1NoAntiquot.formatter := sepByNoAntiquot.formatter -@[combinator_formatter withPosition] def withPosition.formatter (p : Formatter) : Formatter := p -@[combinator_formatter withPositionAfterLinebreak] def withPositionAfterLinebreak.formatter (p : Formatter) : Formatter := p -@[combinator_formatter withoutPosition] def withoutPosition.formatter (p : Formatter) : Formatter := p -@[combinator_formatter withForbidden] def withForbidden.formatter (_tk : Token) (p : Formatter) : Formatter := p -@[combinator_formatter withoutForbidden] def withoutForbidden.formatter (p : Formatter) : Formatter := p @[combinator_formatter withoutInfo] def withoutInfo.formatter (p : Formatter) : Formatter := p -@[combinator_formatter setExpected] -def setExpected.formatter (_expected : List String) (p : Formatter) : Formatter := p - -@[combinator_formatter incQuotDepth] def incQuotDepth.formatter (p : Formatter) : Formatter := p -@[combinator_formatter decQuotDepth] def decQuotDepth.formatter (p : Formatter) : Formatter := p -@[combinator_formatter suppressInsideQuot] def suppressInsideQuot.formatter (p : Formatter) : Formatter := p -@[combinator_formatter evalInsideQuot] def evalInsideQuot.formatter (_declName : Name) (p : Formatter) : Formatter := p - @[combinator_formatter checkWsBefore] def checkWsBefore.formatter : Formatter := do let st ← get if st.leadWord != "" then @@ -501,9 +483,6 @@ def setExpected.formatter (_expected : List String) (p : Formatter) : Formatter @[combinator_formatter pushNone] def pushNone.formatter : Formatter := goLeft -@[combinator_formatter withOpenDecl] def withOpenDecl.formatter (p : Formatter) : Formatter := p -@[combinator_formatter withOpen] def withOpen.formatter (p : Formatter) : Formatter := p - @[combinator_formatter interpolatedStr] def interpolatedStr.formatter (p : Formatter) : Formatter := do visitArgs $ (← getCur).getArgs.reverse.forM fun chunk => @@ -511,8 +490,6 @@ def interpolatedStr.formatter (p : Formatter) : Formatter := do | some str => push str *> goLeft | none => p -@[combinator_formatter dbgTraceState] def dbgTraceState.formatter (_label : String) (p : Formatter) : Formatter := p - @[combinator_formatter _root_.ite, macro_inline] def ite {_ : Type} (c : Prop) [Decidable c] (t e : Formatter) : Formatter := if c then t else e diff --git a/src/Lean/PrettyPrinter/Parenthesizer.lean b/src/Lean/PrettyPrinter/Parenthesizer.lean index 1ff7eae7bd0f..b81ac949daa2 100644 --- a/src/Lean/PrettyPrinter/Parenthesizer.lean +++ b/src/Lean/PrettyPrinter/Parenthesizer.lean @@ -324,12 +324,6 @@ def categoryParser.parenthesizer (cat : Name) (prec : Nat) : Parenthesizer := do -- In this case this node will never be parenthesized since we don't know which parentheses to use. | _ => parenthesizeCategoryCore cat prec -@[combinator_parenthesizer categoryParserOfStack] -def categoryParserOfStack.parenthesizer (offset : Nat) (prec : Nat) : Parenthesizer := do - let st ← get - let stx := st.stxTrav.parents.back.getArg (st.stxTrav.idxs.back - offset) - categoryParser.parenthesizer stx.getId prec - @[combinator_parenthesizer parserOfStack] def parserOfStack.parenthesizer (offset : Nat) (_prec : Nat := 0) : Parenthesizer := do let st ← get @@ -395,6 +389,9 @@ def node.parenthesizer (k : SyntaxNodeKind) (p : Parenthesizer) : Parenthesizer def checkPrec.parenthesizer (prec : Nat) : Parenthesizer := addPrecCheck prec +@[combinator_parenthesizer withFn] +def withFn.parenthesizer (_ : ParserFn → ParserFn) (p : Parenthesizer) : Parenthesizer := p + @[combinator_parenthesizer leadingNode] def leadingNode.parenthesizer (k : SyntaxNodeKind) (prec : Nat) (p : Parenthesizer) : Parenthesizer := do node.parenthesizer k p @@ -470,17 +467,8 @@ def sepByNoAntiquot.parenthesizer (p pSep : Parenthesizer) : Parenthesizer := do @[combinator_parenthesizer withPositionAfterLinebreak] def withPositionAfterLinebreak.parenthesizer (p : Parenthesizer) : Parenthesizer := -- TODO: improve? withPosition.parenthesizer p -@[combinator_parenthesizer withoutPosition] def withoutPosition.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer withForbidden] def withForbidden.parenthesizer (_tk : Parser.Token) (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer withoutForbidden] def withoutForbidden.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer withoutInfo] def withoutInfo.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer setExpected] -def setExpected.parenthesizer (_expected : List String) (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer incQuotDepth] def incQuotDepth.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer decQuotDepth] def decQuotDepth.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer suppressInsideQuot] def suppressInsideQuot.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer evalInsideQuot] def evalInsideQuot.parenthesizer (_declName : Name) (p : Parenthesizer) : Parenthesizer := p +@[combinator_parenthesizer withoutInfo] def withoutInfo.parenthesizer (p : Parenthesizer) : Parenthesizer := p @[combinator_parenthesizer checkStackTop] def checkStackTop.parenthesizer : Parenthesizer := pure () @[combinator_parenthesizer checkWsBefore] def checkWsBefore.parenthesizer : Parenthesizer := pure () @@ -497,9 +485,6 @@ def setExpected.parenthesizer (_expected : List String) (p : Parenthesizer) : Pa @[combinator_parenthesizer pushNone] def pushNone.parenthesizer : Parenthesizer := goLeft -@[combinator_parenthesizer withOpenDecl] def withOpenDecl.parenthesizer (p : Parenthesizer) : Parenthesizer := p -@[combinator_parenthesizer withOpen] def withOpen.parenthesizer (p : Parenthesizer) : Parenthesizer := p - @[combinator_parenthesizer interpolatedStr] def interpolatedStr.parenthesizer (p : Parenthesizer) : Parenthesizer := do visitArgs $ (← getCur).getArgs.reverse.forM fun chunk => @@ -508,8 +493,6 @@ def interpolatedStr.parenthesizer (p : Parenthesizer) : Parenthesizer := do else p -@[combinator_parenthesizer dbgTraceState] def dbgTraceState.parenthesizer (_label : String) (p : Parenthesizer) : Parenthesizer := p - @[combinator_parenthesizer _root_.ite, macro_inline] def ite {_ : Type} (c : Prop) [Decidable c] (t e : Parenthesizer) : Parenthesizer := if c then t else e diff --git a/stage0/src/Lean/Parser/Basic.lean b/stage0/src/Lean/Parser/Basic.lean index 1d393e94369b..f99fb7560e1b 100644 --- a/stage0/src/Lean/Parser/Basic.lean +++ b/stage0/src/Lean/Parser/Basic.lean @@ -89,12 +89,30 @@ structure TokenCacheEntry where stopPos : String.Pos := 0 token : Syntax := Syntax.missing +structure CategoryCacheKey where + cat : Name + prec : Nat + pos : String.Pos + deriving BEq, Hashable + +structure Error where + unexpected : String := "" + expected : List String := [] + deriving Inhabited, BEq + +structure CategoryCacheEntry where + stx : Syntax + lhsPrec : Nat + newPos : String.Pos + errorMsg : Option Error + structure ParserCache where - tokenCache : TokenCacheEntry + tokenCache : TokenCacheEntry + categoryCache : HashMap CategoryCacheKey CategoryCacheEntry -def initCacheForInput (input : String) : ParserCache := { - tokenCache := { startPos := input.endPos + ' ' /- make sure it is not a valid position -/} -} +def initCacheForInput (input : String) : ParserCache where + tokenCache := { startPos := input.endPos + ' ' /- make sure it is not a valid position -/ } + categoryCache := {} abbrev TokenTable := Trie Token @@ -117,7 +135,7 @@ structure ParserModuleContext where env : Environment options : Options -- for name lookup - currNamespace : Name := Name.anonymous + currNamespace : Name := .anonymous openDecls : List OpenDecl := [] structure ParserContext extends InputContext, ParserModuleContext where @@ -129,11 +147,6 @@ structure ParserContext extends InputContext, ParserModuleContext where savedPos? : Option String.Pos := none forbiddenTk? : Option Token := none -structure Error where - unexpected : String := "" - expected : List String := [] - deriving Inhabited, BEq - namespace Error private def expectedToString : List String → String @@ -150,7 +163,8 @@ protected def toString (e : Error) : String := ["expected " ++ expectedToString expected] "; ".intercalate $ unexpected ++ expected -instance : ToString Error := ⟨Error.toString⟩ +instance : ToString Error where + toString := Error.toString def merge (e₁ e₂ : Error) : Error := match e₂ with @@ -199,6 +213,9 @@ def shrinkStack (s : ParserState) (iniStackSz : Nat) : ParserState := def next (s : ParserState) (input : String) (pos : String.Pos) : ParserState := { s with pos := input.next pos } +def next' (s : ParserState) (input : String) (pos : String.Pos) (h : ¬ input.atEnd pos): ParserState := + { s with pos := input.next' pos h } + def toErrorMsg (ctx : ParserContext) (s : ParserState) : String := match s.errorMsg with | none => "" @@ -303,7 +320,8 @@ def toStr : FirstTokens → String | tokens tks => toString tks | optTokens tks => "?" ++ toString tks -instance : ToString FirstTokens := ⟨toStr⟩ +instance : ToString FirstTokens where + toString := toStr end FirstTokens @@ -414,7 +432,7 @@ def checkPrecFn (prec : Nat) : ParserFn := fun c s => else s.mkUnexpectedError "unexpected token at this precedence level; consider parenthesizing the term" def checkPrec (prec : Nat) : Parser := { - info := epsilonInfo, + info := epsilonInfo fn := checkPrecFn prec } @@ -424,7 +442,7 @@ def checkLhsPrecFn (prec : Nat) : ParserFn := fun _ s => else s.mkUnexpectedError "unexpected token at this precedence level; consider parenthesizing the term" def checkLhsPrec (prec : Nat) : Parser := { - info := epsilonInfo, + info := epsilonInfo fn := checkLhsPrecFn prec } @@ -433,7 +451,7 @@ def setLhsPrecFn (prec : Nat) : ParserFn := fun _ s => else { s with lhsPrec := prec } def setLhsPrec (prec : Nat) : Parser := { - info := epsilonInfo, + info := epsilonInfo fn := setLhsPrecFn prec } @@ -441,12 +459,12 @@ private def addQuotDepthFn (i : Int) (p : ParserFn) : ParserFn := fun c s => p { c with quotDepth := c.quotDepth + i |>.toNat } s def incQuotDepth (p : Parser) : Parser := { - info := p.info, + info := p.info fn := addQuotDepthFn 1 p.fn } def decQuotDepth (p : Parser) : Parser := { - info := p.info, + info := p.info fn := addQuotDepthFn (-1) p.fn } @@ -454,7 +472,7 @@ def suppressInsideQuotFn (p : ParserFn) : ParserFn := fun c s => p { c with suppressInsideQuot := true } s def suppressInsideQuot (p : Parser) : Parser := { - info := p.info, + info := p.info fn := suppressInsideQuotFn p.fn } @@ -462,7 +480,7 @@ def leadingNode (n : SyntaxNodeKind) (prec : Nat) (p : Parser) : Parser := checkPrec prec >> node n p >> setLhsPrec prec def trailingNodeAux (n : SyntaxNodeKind) (p : Parser) : TrailingParser := { - info := nodeInfo n p.info, + info := nodeInfo n p.info fn := trailingNodeFn n p.fn } @@ -516,8 +534,8 @@ def orelseFn (p q : ParserFn) : ParserFn := orelseFnCore p q @[noinline] def orelseInfo (p q : ParserInfo) : ParserInfo := { - collectTokens := p.collectTokens ∘ q.collectTokens, - collectKinds := p.collectKinds ∘ q.collectKinds, + collectTokens := p.collectTokens ∘ q.collectTokens + collectKinds := p.collectKinds ∘ q.collectKinds firstTokens := p.firstTokens.merge q.firstTokens } @@ -528,7 +546,7 @@ def orelseFn (p q : ParserFn) : ParserFn := producing a single node kind. Nested `orelse` calls are flattened for this, i.e. `(node k1 p1 <|> node k2 p2) <|> ...` is fine as well. -/ def orelse (p q : Parser) : Parser := { - info := orelseInfo p.info q.info, + info := orelseInfo p.info q.info fn := orelseFn p.fn q.fn } @@ -536,7 +554,7 @@ instance : OrElse Parser where orElse a b := orelse a (b ()) @[noinline] def noFirstTokenInfo (info : ParserInfo) : ParserInfo := { - collectTokens := info.collectTokens, + collectTokens := info.collectTokens collectKinds := info.collectKinds } @@ -547,7 +565,7 @@ def atomicFn (p : ParserFn) : ParserFn := fun c s => | other => other def atomic (p : Parser) : Parser := { - info := p.info, + info := p.info fn := atomicFn p.fn } @@ -559,13 +577,13 @@ def optionalFn (p : ParserFn) : ParserFn := fun c s => s.mkNode nullKind iniSz @[noinline] def optionaInfo (p : ParserInfo) : ParserInfo := { - collectTokens := p.collectTokens, - collectKinds := p.collectKinds, + collectTokens := p.collectTokens + collectKinds := p.collectKinds firstTokens := p.firstTokens.toOptional } def optionalNoAntiquot (p : Parser) : Parser := { - info := optionaInfo p.info, + info := optionaInfo p.info fn := optionalFn p.fn } @@ -576,7 +594,7 @@ def lookaheadFn (p : ParserFn) : ParserFn := fun c s => if s.hasError then s else s.restore iniSz iniPos def lookahead (p : Parser) : Parser := { - info := p.info, + info := p.info fn := lookaheadFn p.fn } @@ -612,7 +630,7 @@ def manyFn (p : ParserFn) : ParserFn := fun c s => s.mkNode nullKind iniSz def manyNoAntiquot (p : Parser) : Parser := { - info := noFirstTokenInfo p.info, + info := noFirstTokenInfo p.info fn := manyFn p.fn } @@ -622,7 +640,7 @@ def many1Fn (p : ParserFn) : ParserFn := fun c s => s.mkNode nullKind iniSz def many1NoAntiquot (p : Parser) : Parser := { - info := p.info, + info := p.info fn := many1Fn p.fn } @@ -663,23 +681,23 @@ def sepBy1Fn (allowTrailingSep : Bool) (p : ParserFn) (sep : ParserFn) : ParserF sepByFnAux p sep allowTrailingSep iniSz false c s @[noinline] def sepByInfo (p sep : ParserInfo) : ParserInfo := { - collectTokens := p.collectTokens ∘ sep.collectTokens, + collectTokens := p.collectTokens ∘ sep.collectTokens collectKinds := p.collectKinds ∘ sep.collectKinds } @[noinline] def sepBy1Info (p sep : ParserInfo) : ParserInfo := { - collectTokens := p.collectTokens ∘ sep.collectTokens, - collectKinds := p.collectKinds ∘ sep.collectKinds, + collectTokens := p.collectTokens ∘ sep.collectTokens + collectKinds := p.collectKinds ∘ sep.collectKinds firstTokens := p.firstTokens } def sepByNoAntiquot (p sep : Parser) (allowTrailingSep : Bool := false) : Parser := { - info := sepByInfo p.info sep.info, + info := sepByInfo p.info sep.info fn := sepByFn allowTrailingSep p.fn sep.fn } def sepBy1NoAntiquot (p sep : Parser) (allowTrailingSep : Bool := false) : Parser := { - info := sepBy1Info p.info sep.info, + info := sepBy1Info p.info sep.info fn := sepBy1Fn allowTrailingSep p.fn sep.fn } @@ -692,12 +710,12 @@ def withResultOfFn (p : ParserFn) (f : Syntax → Syntax) : ParserFn := fun c s s.popSyntax.pushSyntax (f stx) @[noinline] def withResultOfInfo (p : ParserInfo) : ParserInfo := { - collectTokens := p.collectTokens, + collectTokens := p.collectTokens collectKinds := p.collectKinds } def withResultOf (p : Parser) (f : Syntax → Syntax) : Parser := { - info := withResultOfInfo p.info, + info := withResultOfInfo p.info fn := withResultOfFn p.fn f } @@ -706,15 +724,15 @@ def many1Unbox (p : Parser) : Parser := partial def satisfyFn (p : Char → Bool) (errorMsg : String := "unexpected character") : ParserFn := fun c s => let i := s.pos - if c.input.atEnd i then s.mkEOIError - else if p (c.input.get i) then s.next c.input i + if h : c.input.atEnd i then s.mkEOIError + else if p (c.input.get' i h) then s.next' c.input i h else s.mkUnexpectedError errorMsg partial def takeUntilFn (p : Char → Bool) : ParserFn := fun c s => let i := s.pos - if c.input.atEnd i then s - else if p (c.input.get i) then s - else takeUntilFn p c (s.next c.input i) + if h : c.input.atEnd i then s + else if p (c.input.get' i h) then s + else takeUntilFn p c (s.next' c.input i h) def takeWhileFn (p : Char → Bool) : ParserFn := takeUntilFn (fun c => !p c) @@ -725,24 +743,24 @@ def takeWhile1Fn (p : Char → Bool) (errorMsg : String) : ParserFn := partial def finishCommentBlock (nesting : Nat) : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then eoi s + if h : input.atEnd i then eoi s else - let curr := input.get i - let i := input.next i + let curr := input.get' i h + let i := input.next' i h if curr == '-' then - if input.atEnd i then eoi s + if h : input.atEnd i then eoi s else - let curr := input.get i + let curr := input.get' i h if curr == '/' then -- "-/" end of comment - if nesting == 1 then s.next input i - else finishCommentBlock (nesting-1) c (s.next input i) + if nesting == 1 then s.next' input i h + else finishCommentBlock (nesting-1) c (s.next' input i h) else finishCommentBlock nesting c (s.next input i) else if curr == '/' then - if input.atEnd i then eoi s + if h : input.atEnd i then eoi s else - let curr := input.get i - if curr == '-' then finishCommentBlock (nesting+1) c (s.next input i) + let curr := input.get' i h + if curr == '-' then finishCommentBlock (nesting+1) c (s.next' input i h) else finishCommentBlock nesting c (s.setPos i) else finishCommentBlock nesting c (s.setPos i) where @@ -752,19 +770,19 @@ where partial def whitespace : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then s + if h : input.atEnd i then s else - let curr := input.get i + let curr := input.get' i h if curr == '\t' then s.mkUnexpectedError "tabs are not allowed; please configure your editor to expand them" - else if curr.isWhitespace then whitespace c (s.next input i) + else if curr.isWhitespace then whitespace c (s.next' input i h) else if curr == '-' then - let i := input.next i + let i := input.next' i h let curr := input.get i if curr == '-' then andthenFn (takeUntilFn (fun c => c = '\n')) whitespace c (s.next input i) else s else if curr == '/' then - let i := input.next i + let i := input.next' i h let curr := input.get i if curr == '-' then let i := input.next i @@ -774,8 +792,9 @@ partial def whitespace : ParserFn := fun c s => else s else s -def mkEmptySubstringAt (s : String) (p : String.Pos) : Substring := - { str := s, startPos := p, stopPos := p } +def mkEmptySubstringAt (s : String) (p : String.Pos) : Substring := { + str := s, startPos := p, stopPos := p +} private def rawAux (startPos : String.Pos) (trailingWs : Bool) : ParserFn := fun c s => let input := c.input @@ -802,31 +821,32 @@ def rawFn (p : ParserFn) (trailingWs := false) : ParserFn := fun c s => def chFn (c : Char) (trailingWs := false) : ParserFn := rawFn (satisfyFn (fun d => c == d) ("'" ++ toString c ++ "'")) trailingWs -def rawCh (c : Char) (trailingWs := false) : Parser := - { fn := chFn c trailingWs } +def rawCh (c : Char) (trailingWs := false) : Parser := { + fn := chFn c trailingWs +} def hexDigitFn : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then s.mkEOIError + if h : input.atEnd i then s.mkEOIError else - let curr := input.get i - let i := input.next i + let curr := input.get' i h + let i := input.next' i h if curr.isDigit || ('a' <= curr && curr <= 'f') || ('A' <= curr && curr <= 'F') then s.setPos i else s.mkUnexpectedError "invalid hexadecimal numeral" def quotedCharCoreFn (isQuotable : Char → Bool) : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then s.mkEOIError + if h : input.atEnd i then s.mkEOIError else - let curr := input.get i + let curr := input.get' i h if isQuotable curr then - s.next input i + s.next' input i h else if curr == 'x' then - andthenFn hexDigitFn hexDigitFn c (s.next input i) + andthenFn hexDigitFn hexDigitFn c (s.next' input i h) else if curr == 'u' then - andthenFn hexDigitFn (andthenFn hexDigitFn (andthenFn hexDigitFn hexDigitFn)) c (s.next input i) + andthenFn hexDigitFn (andthenFn hexDigitFn (andthenFn hexDigitFn hexDigitFn)) c (s.next' input i h) else s.mkUnexpectedError "invalid escape sequence" @@ -851,10 +871,10 @@ def mkNodeToken (n : SyntaxNodeKind) (startPos : String.Pos) : ParserFn := fun c def charLitFnAux (startPos : String.Pos) : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then s.mkEOIError + if h : input.atEnd i then s.mkEOIError else - let curr := input.get i - let s := s.setPos (input.next i) + let curr := input.get' i h + let s := s.setPos (input.next' i h) let s := if curr == '\\' then quotedCharFn c s else s if s.hasError then s else @@ -867,10 +887,10 @@ def charLitFnAux (startPos : String.Pos) : ParserFn := fun c s => partial def strLitFnAux (startPos : String.Pos) : ParserFn := fun c s => let input := c.input let i := s.pos - if input.atEnd i then s.mkUnexpectedErrorAt "unterminated string literal" startPos + if h : input.atEnd i then s.mkUnexpectedErrorAt "unterminated string literal" startPos else - let curr := input.get i - let s := s.setPos (input.next i) + let curr := input.get' i h + let s := s.setPos (input.next' i h) if curr == '\"' then mkNodeToken strLitKind startPos c s else if curr == '\\' then andthenFn quotedCharFn (strLitFnAux startPos) c s @@ -932,11 +952,11 @@ def hexNumberFn (startPos : String.Pos) : ParserFn := fun c s => def numberFnAux : ParserFn := fun c s => let input := c.input let startPos := s.pos - if input.atEnd startPos then s.mkEOIError + if h : input.atEnd startPos then s.mkEOIError else - let curr := input.get startPos + let curr := input.get' startPos h if curr == '0' then - let i := input.next startPos + let i := input.next' startPos h let curr := input.get i if curr == 'b' || curr == 'B' then binNumberFn startPos c (s.next input i) @@ -1006,30 +1026,32 @@ def mkIdResult (startPos : String.Pos) (tk : Option Token) (val : Name) : Parser s.pushSyntax atom partial def identFnAux (startPos : String.Pos) (tk : Option Token) (r : Name) : ParserFn := - let rec parse (r : Name) (c s) := Id.run do + let rec parse (r : Name) (c s) := let input := c.input let i := s.pos - if input.atEnd i then - return s.mkEOIError - let curr := input.get i - if isIdBeginEscape curr then - let startPart := input.next i - let s := takeUntilFn isIdEndEscape c (s.setPos startPart) - if input.atEnd s.pos then - return s.mkUnexpectedErrorAt "unterminated identifier escape" startPart - let stopPart := s.pos - let s := s.next c.input s.pos - let r := Name.mkStr r (input.extract startPart stopPart) - if isIdCont input s then - let s := s.next input s.pos - parse r c s - else - mkIdResult startPos tk r c s + if h : input.atEnd i then + s.mkEOIError + else + let curr := input.get' i h + if isIdBeginEscape curr then + let startPart := input.next' i h + let s := takeUntilFn isIdEndEscape c (s.setPos startPart) + if h : input.atEnd s.pos then + s.mkUnexpectedErrorAt "unterminated identifier escape" startPart + else + let stopPart := s.pos + let s := s.next' c.input s.pos h + let r := .str r (input.extract startPart stopPart) + if isIdCont input s then + let s := s.next input s.pos + parse r c s + else + mkIdResult startPos tk r c s else if isIdFirst curr then let startPart := i let s := takeWhileFn isIdRest c (s.next input i) let stopPart := s.pos - let r := Name.mkStr r (input.extract startPart stopPart) + let r := .str r (input.extract startPart stopPart) if isIdCont input s then let s := s.next input s.pos parse r c s @@ -1044,13 +1066,13 @@ private def isIdFirstOrBeginEscape (c : Char) : Bool := private def nameLitAux (startPos : String.Pos) : ParserFn := fun c s => let input := c.input - let s := identFnAux startPos none Name.anonymous c (s.next input startPos) + let s := identFnAux startPos none .anonymous c (s.next input startPos) if s.hasError then s else let stx := s.stxStack.back match stx with - | Syntax.ident info rawStr _ _ => + | .ident info rawStr _ _ => let s := s.popSyntax s.pushSyntax (Syntax.mkNameLit rawStr.toString info) | _ => s.mkError "invalid Name literal" @@ -1069,16 +1091,16 @@ private def tokenFnAux : ParserFn := fun c s => nameLitAux i c s else let (_, tk) := c.tokens.matchPrefix input i - identFnAux i tk Name.anonymous c s + identFnAux i tk .anonymous c s -private def updateCache (startPos : String.Pos) (s : ParserState) : ParserState := +private def updateTokenCache (startPos : String.Pos) (s : ParserState) : ParserState := -- do not cache token parsing errors, which are rare and usually fatal and thus not worth an extra field in `TokenCache` match s with - | ⟨stack, lhsPrec, pos, _, none⟩ => + | ⟨stack, lhsPrec, pos, ⟨_, catCache⟩, none⟩ => if stack.size == 0 then s else let tk := stack.back - ⟨stack, lhsPrec, pos, { tokenCache := { startPos := startPos, stopPos := pos, token := tk } }, none⟩ + ⟨stack, lhsPrec, pos, ⟨{ startPos := startPos, stopPos := pos, token := tk }, catCache⟩, none⟩ | other => other def tokenFn (expected : List String := []) : ParserFn := fun c s => @@ -1092,21 +1114,21 @@ def tokenFn (expected : List String := []) : ParserFn := fun c s => s.setPos tkc.stopPos else let s := tokenFnAux c s - updateCache i s + updateTokenCache i s def peekTokenAux (c : ParserContext) (s : ParserState) : ParserState × Except ParserState Syntax := let iniSz := s.stackSize let iniPos := s.pos let s := tokenFn [] c s - if let some _ := s.errorMsg then (s.restore iniSz iniPos, Except.error s) + if let some _ := s.errorMsg then (s.restore iniSz iniPos, .error s) else let stx := s.stxStack.back - (s.restore iniSz iniPos, Except.ok stx) + (s.restore iniSz iniPos, .ok stx) def peekToken (c : ParserContext) (s : ParserState) : ParserState × Except ParserState Syntax := let tkc := s.cache.tokenCache if tkc.startPos == s.pos then - (s, Except.ok tkc.token) + (s, .ok tkc.token) else peekTokenAux c s @@ -1115,7 +1137,7 @@ def rawIdentFn : ParserFn := fun c s => let input := c.input let i := s.pos if input.atEnd i then s.mkEOIError - else identFnAux i none Name.anonymous c s + else identFnAux i none .anonymous c s def satisfySymbolFn (p : String → Bool) (expected : List String) : ParserFn := fun c s => let initStackSz := s.stackSize @@ -1125,14 +1147,14 @@ def satisfySymbolFn (p : String → Bool) (expected : List String) : ParserFn := s else match s.stxStack.back with - | Syntax.atom _ sym => if p sym then s else s.mkErrorsAt expected startPos initStackSz - | _ => s.mkErrorsAt expected startPos initStackSz + | .atom _ sym => if p sym then s else s.mkErrorsAt expected startPos initStackSz + | _ => s.mkErrorsAt expected startPos initStackSz def symbolFnAux (sym : String) (errorMsg : String) : ParserFn := satisfySymbolFn (fun s => s == sym) [errorMsg] def symbolInfo (sym : String) : ParserInfo := { - collectTokens := fun tks => sym :: tks, + collectTokens := fun tks => sym :: tks firstTokens := FirstTokens.tokens [ sym ] } @@ -1141,13 +1163,13 @@ def symbolFn (sym : String) : ParserFn := def symbolNoAntiquot (sym : String) : Parser := let sym := sym.trim - { info := symbolInfo sym, + { info := symbolInfo sym fn := symbolFn sym } def checkTailNoWs (prev : Syntax) : Bool := match prev.getTailInfo with - | SourceInfo.original _ _ trailing _ => trailing.stopPos == trailing.startPos - | _ => false + | .original _ _ trailing _ => trailing.stopPos == trailing.startPos + | _ => false /-- Check if the following token is the symbol _or_ identifier `sym`. Useful for parsing local tokens that have not been added to the token table (but may have @@ -1163,9 +1185,9 @@ def nonReservedSymbolFnAux (sym : String) (errorMsg : String) : ParserFn := fun if s.hasError then s else match s.stxStack.back with - | Syntax.atom _ sym' => + | .atom _ sym' => if sym == sym' then s else s.mkErrorAt errorMsg startPos initStackSz - | Syntax.ident info rawVal _ _ => + | .ident info rawVal _ _ => if sym == rawVal.toString then let s := s.popSyntax s.pushSyntax (Syntax.atom info sym) @@ -1179,9 +1201,9 @@ def nonReservedSymbolFn (sym : String) : ParserFn := def nonReservedSymbolInfo (sym : String) (includeIdent : Bool) : ParserInfo := { firstTokens := if includeIdent then - FirstTokens.tokens [ sym, "ident" ] + .tokens [ sym, "ident" ] else - FirstTokens.tokens [ sym ] + .tokens [ sym ] } def nonReservedSymbolNoAntiquot (sym : String) (includeIdent := false) : Parser := @@ -1191,31 +1213,32 @@ def nonReservedSymbolNoAntiquot (sym : String) (includeIdent := false) : Parser partial def strAux (sym : String) (errorMsg : String) (j : String.Pos) :ParserFn := let rec parse (j c s) := - if sym.atEnd j then s + if h₁ : sym.atEnd j then s else let i := s.pos let input := c.input - if input.atEnd i || sym.get j != input.get i then s.mkError errorMsg - else parse (sym.next j) c (s.next input i) + if h₂ : input.atEnd i then s.mkError errorMsg + else if sym.get' j h₁ != input.get' i h₂ then s.mkError errorMsg + else parse (sym.next' j h₁) c (s.next' input i h₂) parse j def checkTailWs (prev : Syntax) : Bool := match prev.getTailInfo with - | SourceInfo.original _ _ trailing _ => trailing.stopPos > trailing.startPos - | _ => false + | .original _ _ trailing _ => trailing.stopPos > trailing.startPos + | _ => false def checkWsBeforeFn (errorMsg : String) : ParserFn := fun _ s => let prev := s.stxStack.back if checkTailWs prev then s else s.mkError errorMsg def checkWsBefore (errorMsg : String := "space before") : Parser := { - info := epsilonInfo, + info := epsilonInfo fn := checkWsBeforeFn errorMsg } def checkTailLinebreak (prev : Syntax) : Bool := match prev.getTailInfo with - | SourceInfo.original _ _ trailing _ => trailing.contains '\n' + | .original _ _ trailing _ => trailing.contains '\n' | _ => false def checkLinebreakBeforeFn (errorMsg : String) : ParserFn := fun _ s => @@ -1237,7 +1260,7 @@ def checkNoWsBeforeFn (errorMsg : String) : ParserFn := fun _ s => if checkTailNoWs prev then s else s.mkError errorMsg def checkNoWsBefore (errorMsg : String := "no space before") : Parser := { - info := epsilonInfo, + info := epsilonInfo fn := checkNoWsBeforeFn errorMsg } @@ -1245,7 +1268,7 @@ def unicodeSymbolFnAux (sym asciiSym : String) (expected : List String) : Parser satisfySymbolFn (fun s => s == sym || s == asciiSym) expected def unicodeSymbolInfo (sym asciiSym : String) : ParserInfo := { - collectTokens := fun tks => sym :: asciiSym :: tks, + collectTokens := fun tks => sym :: asciiSym :: tks firstTokens := FirstTokens.tokens [ sym, asciiSym ] } @@ -1255,33 +1278,31 @@ def unicodeSymbolFn (sym asciiSym : String) : ParserFn := def unicodeSymbolNoAntiquot (sym asciiSym : String) : Parser := let sym := sym.trim let asciiSym := asciiSym.trim - { info := unicodeSymbolInfo sym asciiSym, + { info := unicodeSymbolInfo sym asciiSym fn := unicodeSymbolFn sym asciiSym } def mkAtomicInfo (k : String) : ParserInfo := { firstTokens := FirstTokens.tokens [ k ] } -def numLitFn : ParserFn := - fun c s => - let initStackSz := s.stackSize - let iniPos := s.pos - let s := tokenFn ["numeral"] c s - if !s.hasError && !(s.stxStack.back.isOfKind numLitKind) then s.mkErrorAt "numeral" iniPos initStackSz else s +def numLitFn : ParserFn := fun c s => + let initStackSz := s.stackSize + let iniPos := s.pos + let s := tokenFn ["numeral"] c s + if !s.hasError && !(s.stxStack.back.isOfKind numLitKind) then s.mkErrorAt "numeral" iniPos initStackSz else s def numLitNoAntiquot : Parser := { - fn := numLitFn, + fn := numLitFn info := mkAtomicInfo "num" } -def scientificLitFn : ParserFn := - fun c s => - let initStackSz := s.stackSize - let iniPos := s.pos - let s := tokenFn ["scientific number"] c s - if !s.hasError && !(s.stxStack.back.isOfKind scientificLitKind) then s.mkErrorAt "scientific number" iniPos initStackSz else s +def scientificLitFn : ParserFn := fun c s => + let initStackSz := s.stackSize + let iniPos := s.pos + let s := tokenFn ["scientific number"] c s + if !s.hasError && !(s.stxStack.back.isOfKind scientificLitKind) then s.mkErrorAt "scientific number" iniPos initStackSz else s def scientificLitNoAntiquot : Parser := { - fn := scientificLitFn, + fn := scientificLitFn info := mkAtomicInfo "scientific" } @@ -1292,7 +1313,7 @@ def strLitFn : ParserFn := fun c s => if !s.hasError && !(s.stxStack.back.isOfKind strLitKind) then s.mkErrorAt "string literal" iniPos initStackSz else s def strLitNoAntiquot : Parser := { - fn := strLitFn, + fn := strLitFn info := mkAtomicInfo "str" } @@ -1303,7 +1324,7 @@ def charLitFn : ParserFn := fun c s => if !s.hasError && !(s.stxStack.back.isOfKind charLitKind) then s.mkErrorAt "character literal" iniPos initStackSz else s def charLitNoAntiquot : Parser := { - fn := charLitFn, + fn := charLitFn info := mkAtomicInfo "char" } @@ -1314,7 +1335,7 @@ def nameLitFn : ParserFn := fun c s => if !s.hasError && !(s.stxStack.back.isOfKind nameLitKind) then s.mkErrorAt "Name literal" iniPos initStackSz else s def nameLitNoAntiquot : Parser := { - fn := nameLitFn, + fn := nameLitFn info := mkAtomicInfo "name" } @@ -1325,7 +1346,7 @@ def identFn : ParserFn := fun c s => if !s.hasError && !(s.stxStack.back.isIdent) then s.mkErrorAt "identifier" iniPos initStackSz else s def identNoAntiquot : Parser := { - fn := identFn, + fn := identFn info := mkAtomicInfo "ident" } @@ -1340,11 +1361,11 @@ def identEqFn (id : Name) : ParserFn := fun c s => if s.hasError then s else match s.stxStack.back with - | Syntax.ident _ _ val _ => if val != id then s.mkErrorAt ("expected identifier '" ++ toString id ++ "'") iniPos initStackSz else s + | .ident _ _ val _ => if val != id then s.mkErrorAt ("expected identifier '" ++ toString id ++ "'") iniPos initStackSz else s | _ => s.mkErrorAt "identifier" iniPos initStackSz def identEq (id : Name) : Parser := { - fn := identEqFn id, + fn := identEqFn id info := mkAtomicInfo "ident" } @@ -1543,8 +1564,9 @@ def eoiFn : ParserFn := fun c s => if c.input.atEnd i then s else s.mkError "expected end of file" -def eoi : Parser := - { fn := eoiFn } +def eoi : Parser := { + fn := eoiFn +} /-- A multimap indexed by tokens. Used for indexing parsers by their leading token. -/ def TokenMap (α : Type) := RBMap Name (List α) Name.quickCmp @@ -1553,10 +1575,11 @@ namespace TokenMap def insert (map : TokenMap α) (k : Name) (v : α) : TokenMap α := match map.find? k with - | none => RBMap.insert map k [v] - | some vs => RBMap.insert map k (v::vs) + | none => .insert map k [v] + | some vs => .insert map k (v::vs) -instance : Inhabited (TokenMap α) := ⟨RBMap.empty⟩ +instance : Inhabited (TokenMap α) where + default := RBMap.empty instance : EmptyCollection (TokenMap α) := ⟨RBMap.empty⟩ @@ -1570,7 +1593,8 @@ structure PrattParsingTables where trailingTable : TokenMap (Parser × Nat) := {} trailingParsers : List (Parser × Nat) := [] -- for supporting parsers such as function application -instance : Inhabited PrattParsingTables := ⟨{}⟩ +instance : Inhabited PrattParsingTables where + default := {} /-- The type `LeadingIdentBehavior` specifies how the parsing table @@ -1644,15 +1668,15 @@ def indexed {α : Type} (map : TokenMap α) (c : ParserContext) (s : ParserState | some as => (s, as) | _ => (s, []) match stx with - | Except.ok (Syntax.atom _ sym) => find (Name.mkSimple sym) - | Except.ok (Syntax.ident _ _ val _) => + | .ok (.atom _ sym) => find (.mkSimple sym) + | .ok (.ident _ _ val _) => match behavior with - | LeadingIdentBehavior.default => find identKind - | LeadingIdentBehavior.symbol => + | .default => find identKind + | .symbol => match map.find? val with | some as => (s, as) | none => find identKind - | LeadingIdentBehavior.both => + | .both => match map.find? val with | some as => if val == identKind then @@ -1662,9 +1686,9 @@ def indexed {α : Type} (map : TokenMap α) (c : ParserContext) (s : ParserState | some as' => (s, as ++ as') | _ => (s, as) | none => find identKind - | Except.ok (Syntax.node _ k _) => find k - | Except.ok _ => (s, []) - | Except.error s' => (s', []) + | .ok (.node _ k _) => find k + | .ok _ => (s, []) + | .error s' => (s', []) abbrev CategoryParserFn := Name → ParserFn @@ -1676,7 +1700,17 @@ def categoryParserFn (catName : Name) : ParserFn := fun ctx s => categoryParserFnExtension.getState ctx.env catName ctx s def categoryParser (catName : Name) (prec : Nat) : Parser := { - fn := fun c s => categoryParserFn catName { c with prec := prec } s + fn := fun c s => Id.run do + let key := ⟨catName, prec, s.pos⟩ + if let some r := s.cache.categoryCache.find? key then + match s with + | ⟨stack, _, _, cache, _⟩ => return ⟨stack.push r.stx, r.lhsPrec, r.newPos, cache, r.errorMsg⟩ + let initStackSz := s.stackSize + let s := categoryParserFn catName { c with prec := prec } s + if s.stackSize > initStackSz + 1 then + panic! "categoryParser: unexpected stack growth" + let s := if s.stackSize == initStackSz then s.pushSyntax .missing else s + { s with cache.categoryCache := s.cache.categoryCache.insert key ⟨s.stxStack.back, s.lhsPrec, s.pos, s.errorMsg⟩ } } -- Define `termParser` here because we need it for antiquotations @@ -1694,9 +1728,9 @@ def checkNoImmediateColon : Parser := { if checkTailNoWs prev then let input := c.input let i := s.pos - if input.atEnd i then s + if h : input.atEnd i then s else - let curr := input.get i + let curr := input.get' i h if curr == ':' then s.mkUnexpectedError "unexpected ':'" else s @@ -1708,11 +1742,13 @@ def setExpectedFn (expected : List String) (p : ParserFn) : ParserFn := fun c s | s'@{ errorMsg := some msg, .. } => { s' with errorMsg := some { msg with expected } } | s' => s' -def setExpected (expected : List String) (p : Parser) : Parser := - { fn := setExpectedFn expected p.fn, info := p.info } +def setExpected (expected : List String) (p : Parser) : Parser := { + fn := setExpectedFn expected p.fn, info := p.info +} -def pushNone : Parser := - { fn := fun _ s => s.pushSyntax mkNullNode } +def pushNone : Parser := { + fn := fun _ s => s.pushSyntax mkNullNode +} -- We support three kinds of antiquotations: `$id`, `$_`, and `$(t)`, where `id` is a term identifier and `t` is a term. def antiquotNestedExpr : Parser := node `antiquotNestedExpr (symbolNoAntiquot "(" >> decQuotDepth termParser >> symbolNoAntiquot ")") @@ -1741,7 +1777,8 @@ def tokenWithAntiquot (p : Parser) : Parser where def symbol (sym : String) : Parser := tokenWithAntiquot (symbolNoAntiquot sym) -instance : Coe String Parser := ⟨fun s => symbol s ⟩ +instance : Coe String Parser where + coe := symbol def nonReservedSymbol (sym : String) (includeIdent := false) : Parser := tokenWithAntiquot (nonReservedSymbolNoAntiquot sym includeIdent) @@ -1754,13 +1791,13 @@ def unicodeSymbol (sym asciiSym : String) : Parser := `kind` is embedded in the antiquotation's kind, and checked at syntax `match` unless `isPseudoKind` is true. Antiquotations can be escaped as in `$$e`, which produces the syntax tree for `$e`. -/ def mkAntiquot (name : String) (kind : SyntaxNodeKind) (anonymous := true) (isPseudoKind := false) : Parser := - let kind := kind ++ (if isPseudoKind then `pseudo else Name.anonymous) ++ `antiquot - let nameP := node `antiquotName $ checkNoWsBefore ("no space before ':" ++ name ++ "'") >> symbol ":" >> nonReservedSymbol name + let kind := kind ++ (if isPseudoKind then `pseudo else .anonymous) ++ `antiquot + let nameP := node `antiquotName <| checkNoWsBefore ("no space before ':" ++ name ++ "'") >> symbol ":" >> nonReservedSymbol name -- if parsing the kind fails and `anonymous` is true, check that we're not ignoring a different -- antiquotation kind via `noImmediateColon` let nameP := if anonymous then nameP <|> checkNoImmediateColon >> pushNone else nameP -- antiquotations are not part of the "standard" syntax, so hide "expected '$'" on error - leadingNode kind maxPrec $ atomic $ + leadingNode kind maxPrec <| atomic <| setExpected [] "$" >> manyNoAntiquot (checkNoWsBefore "" >> "$") >> checkNoWsBefore "no space before spliced term" >> antiquotExpr >> @@ -1779,17 +1816,18 @@ def withAntiquotFn (antiquotP p : ParserFn) (isCatAntiquot := false) : ParserFn /-- Optimized version of `mkAntiquot ... <|> p`. -/ def withAntiquot (antiquotP p : Parser) : Parser := { - fn := withAntiquotFn antiquotP.fn p.fn, + fn := withAntiquotFn antiquotP.fn p.fn info := orelseInfo antiquotP.info p.info } -def withoutInfo (p : Parser) : Parser := - { fn := p.fn } +def withoutInfo (p : Parser) : Parser := { + fn := p.fn +} /-- Parse `$[p]suffix`, e.g. `$[p],*`. -/ def mkAntiquotSplice (kind : SyntaxNodeKind) (p suffix : Parser) : Parser := let kind := kind ++ `antiquot_scope - leadingNode kind maxPrec $ atomic $ + leadingNode kind maxPrec <| atomic <| setExpected [] "$" >> manyNoAntiquot (checkNoWsBefore "" >> "$") >> checkNoWsBefore "no space before spliced term" >> symbol "[" >> node nullKind p >> symbol "]" >> @@ -1836,15 +1874,20 @@ def sepBy1 (p : Parser) (sep : String) (psep : Parser := symbol sep) (allowTrail def categoryParserOfStackFn (offset : Nat) : ParserFn := fun ctx s => let stack := s.stxStack - if stack.size < offset + 1 then + if h : stack.size < offset + 1 then s.mkUnexpectedError ("failed to determine parser category using syntax stack, stack is too small") else - match stack.get! (stack.size - offset - 1) with - | Syntax.ident _ _ catName _ => categoryParserFn catName ctx s + have : stack.size - (offset + 1) < stack.size := by + apply Nat.sub_lt <;> simp_all_arith + apply Nat.le_trans (m := offset + 1) + apply Nat.le_add_left; assumption + match stack[stack.size - (offset + 1)] with + | .ident _ _ catName _ => categoryParserFn catName ctx s | _ => s.mkUnexpectedError ("failed to determine parser category using syntax stack, the specified element on the stack is not an identifier") -def categoryParserOfStack (offset : Nat) (prec : Nat := 0) : Parser := - { fn := fun c s => categoryParserOfStackFn offset { c with prec := prec } s } +def categoryParserOfStack (offset : Nat) (prec : Nat := 0) : Parser := { + fn := fun c s => categoryParserOfStackFn offset { c with prec := prec } s +} private def mkResult (s : ParserState) (iniSz : Nat) : ParserState := if s.stackSize == iniSz + 1 then s @@ -1927,12 +1970,12 @@ def fieldIdxFn : ParserFn := fun c s => def fieldIdx : Parser := withAntiquot (mkAntiquot "fieldIdx" `fieldIdx) { - fn := fieldIdxFn, + fn := fieldIdxFn info := mkAtomicInfo "fieldIdx" } def skip : Parser := { - fn := fun _ s => s, + fn := fun _ s => s info := epsilonInfo } @@ -1941,7 +1984,7 @@ end Parser namespace Syntax section -variable {β : Type} {m : Type → Type} [Monad m] +variable [Monad m] def foldArgsM (s : Syntax) (f : Syntax → β → m β) (b : β) : m β := s.getArgs.foldlM (flip f) b diff --git a/stage0/src/runtime/object.cpp b/stage0/src/runtime/object.cpp index 06f7d2284436..555063f46b1e 100644 --- a/stage0/src/runtime/object.cpp +++ b/stage0/src/runtime/object.cpp @@ -1697,7 +1697,6 @@ static bool lean_string_utf8_get_core(char const * str, usize size, usize i, uin unsigned c = static_cast(str[i]); /* zero continuation (0 to 127) */ if ((c & 0x80) == 0) { - i++; result = c; return true; } @@ -1707,7 +1706,6 @@ static bool lean_string_utf8_get_core(char const * str, usize size, usize i, uin unsigned c1 = static_cast(str[i+1]); result = ((c & 0x1f) << 6) | (c1 & 0x3f); if (result >= 128) { - i += 2; return true; } } @@ -1718,7 +1716,6 @@ static bool lean_string_utf8_get_core(char const * str, usize size, usize i, uin unsigned c2 = static_cast(str[i+2]); result = ((c & 0x0f) << 12) | ((c1 & 0x3f) << 6) | (c2 & 0x3f); if (result >= 2048 && (result < 55296 || result > 57343)) { - i += 3; return true; } } @@ -1730,7 +1727,6 @@ static bool lean_string_utf8_get_core(char const * str, usize size, usize i, uin unsigned c3 = static_cast(str[i+3]); result = ((c & 0x07) << 18) | ((c1 & 0x3f) << 12) | ((c2 & 0x3f) << 6) | (c3 & 0x3f); if (result >= 65536 && result <= 1114111) { - i += 4; return true; } } diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index aae6413ccc9c..82742d5e989d 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -6,7 +6,7 @@ options get_default_options() { // see https://leanprover.github.io/lean4/doc/dev/bootstrap.html#further-bootstrapping-complications #if LEAN_IS_STAGE0 == 1 // switch to `true` for ABI-breaking changes affecting meta code - opts = opts.update({"interpreter", "prefer_native"}, false); + opts = opts.update({"interpreter", "prefer_native"}, true); // switch to `true` for changing built-in parsers used in quotations opts = opts.update({"internal", "parseQuotWithCurrentStage"}, false); opts = opts.update({"pp", "rawOnError"}, true); diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index 9b6afca1bf39..6fb88c4da71e 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -214,7 +214,6 @@ static lean_object* l_Lean_Elab_Command_elabModuleDoc___closed__1; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Elab_Command_elabUniverse___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice___closed__5; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth___closed__5; lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -317,6 +316,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Built LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSynth(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNonComputableSection_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabAddDeclDoc___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabVariable___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange___closed__1; @@ -371,7 +371,7 @@ static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEvalUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabInitQuot_declRange___closed__2; static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___closed__20; -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabExport___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabNamespace_declRange___closed__2; lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); @@ -484,6 +484,7 @@ static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___closed__34; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveNamespace___at_Lean_Elab_Command_elabOpen___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabChoice___closed__2; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_containsId___boxed(lean_object*, lean_object*); @@ -632,7 +633,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Built LEAN_EXPORT lean_object* l_Lean_MonadQuotation_addMacroScope___at_Lean_Elab_Command_elabVariable___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabExport___spec__19(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabExport___spec__10___closed__3; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Elab_Command_elabSetOption___spec__1___closed__4; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___lambda__1___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Command_elabInitQuot___closed__5; @@ -905,6 +906,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheckFailure_declRange_ static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___closed__3; static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___closed__27; static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__12; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabReduce_declRange___closed__5; static lean_object* l_Lean_Elab_Command_elabEvalUnsafe___lambda__4___closed__6; static lean_object* l_Lean_setEnv___at_Lean_Elab_Command_elabEvalUnsafe___spec__2___closed__5; @@ -4712,7 +4714,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; @@ -4779,7 +4781,15 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; @@ -4824,13 +4834,21 @@ size_t x_14; size_t x_15; lean_object* x_16; x_14 = 0; x_15 = lean_usize_of_nat(x_8); lean_dec(x_8); -x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(x_2, x_7, x_14, x_15, x_3, x_4, x_5, x_6); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg(x_2, x_7, x_14, x_15, x_3, x_4, x_5, x_6); lean_dec(x_7); return x_16; } } } } +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2___rarg), 6, 0); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___at_Lean_Elab_Command_elabUniverse___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -4846,7 +4864,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_alloc_closure((void*)(l_Lean_Syntax_forArgsM___at_Lean_Elab_Command_elabUniverse___spec__1___lambda__1___boxed), 6, 1); lean_closure_set(x_6, 0, x_2); x_7 = lean_box(0); -x_8 = l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2(x_1, x_6, x_7, x_3, x_4, x_5); +x_8 = l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2___rarg(x_1, x_6, x_7, x_3, x_4, x_5); return x_8; } } @@ -4870,7 +4888,7 @@ x_8 = l_Lean_Syntax_forArgsM___at_Lean_Elab_Command_elabUniverse___spec__1(x_6, return x_8; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { size_t x_9; size_t x_10; lean_object* x_11; @@ -4878,7 +4896,7 @@ x_9 = lean_unbox_usize(x_3); lean_dec(x_3); x_10 = lean_unbox_usize(x_4); lean_dec(x_4); -x_11 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +x_11 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabUniverse___spec__3___rarg(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); lean_dec(x_2); return x_11; } diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index 4c1f1bce2900..5c262623f404 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -559,7 +559,7 @@ uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_ uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___closed__4; lean_object* l_Lean_quoteNameMk(lean_object*); -uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586_(uint8_t, uint8_t); +uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_Command_elabSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkLeftRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__2___rarg(lean_object*); @@ -5009,7 +5009,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1( lean_object* x_12; uint8_t x_54; uint8_t x_55; uint8_t x_56; x_54 = lean_ctor_get_uint8(x_3, sizeof(void*)*1 + 2); x_55 = 0; -x_56 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586_(x_54, x_55); +x_56 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950_(x_54, x_55); if (x_56 == 0) { uint8_t x_57; diff --git a/stage0/stdlib/Lean/Parser/Basic.c b/stage0/stdlib/Lean/Parser/Basic.c index 85cc7acde18f..82d624089972 100644 --- a/stage0/stdlib/Lean/Parser/Basic.c +++ b/stage0/stdlib/Lean/Parser/Basic.c @@ -22,15 +22,16 @@ LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__2___boxed(lean LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserOfStack___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Parser_Error_toString___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Parser_info___default___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_finishCommentBlock(lean_object*, lean_object*, lean_object*); +size_t l___private_Lean_Data_HashMap_0__Lean_HashMapImp_mkIdx(lean_object*, uint64_t, lean_object*); static lean_object* l_Lean_Parser_categoryParserFn___closed__1; lean_object* l_String_csize(uint32_t); lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedError; LEAN_EXPORT lean_object* l_Lean_Parser_andthenInfo___elambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_next_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_error___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withResultOf___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux(lean_object*, lean_object*, lean_object*); @@ -55,13 +56,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_prattParser(lean_object*, lean_object*, u LEAN_EXPORT lean_object* l_Lean_Parser_andthenInfo(lean_object*, lean_object*); extern lean_object* l_String_instInhabitedString; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkError(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn_parseOptDot(lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Parser_Error_toString___spec__1___closed__1; static lean_object* l_Lean_Parser_strLitNoAntiquot___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByInfo(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Parser_info___default; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_shrinkStack___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -69,7 +70,6 @@ lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Parser_eoi___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_eoi___closed__1; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_manyNoAntiquot(lean_object*); static lean_object* l_Lean_Parser_notFollowedByFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotSuffixSplice___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,6 +92,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEq___elambda__1(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Parser_eoi; LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrecFn(lean_object*, lean_object*, lean_object*); extern uint32_t l_Lean_idBeginEscape; +LEAN_EXPORT lean_object* l_Lean_Parser_instBEqCategoryCacheKey; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_errorMsg___default; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optionaInfo___elambda__2(lean_object*, lean_object*); @@ -99,6 +100,7 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_binNumberFn___closed__1; static lean_object* l_Lean_Parser_hexNumberFn___closed__1; LEAN_EXPORT uint8_t l_Lean_Parser_unicodeSymbolFnAux___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_instBEqCategoryCacheKey___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_numberFnAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withoutInfo(lean_object*); @@ -108,19 +110,21 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__5__ extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_Parser_lookahead(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolInfo___elambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_skip___closed__2; +static lean_object* l_Lean_Parser_Error_unexpected___default___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__6___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_isQuotableCharDefault(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_errorAtSavedPos___closed__1; +static lean_object* l_Lean_Parser_instCoeStringParser___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_nameLitNoAntiquot; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_runLongestMatchParser___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserInfo_collectKinds___default___boxed(lean_object*); static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_whitespace___lambda__1___boxed(lean_object*); +lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Parser_instInhabitedError___closed__1; lean_object* l_Lean_Parser_Trie_matchPrefix___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_stackSize___boxed(lean_object*); @@ -151,6 +155,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserInfo_collectKinds___default(lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_many1NoAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_leadingTable___default; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_pushSyntax(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn___lambda__1___boxed(lean_object*, lean_object*); @@ -174,13 +179,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolFn(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_error___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotSuffixSplice___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Char_isDigit(uint32_t); -LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkNode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_binNumberFn___lambda__1(uint32_t); static lean_object* l_Lean_Parser_instInhabitedParserInfo___closed__1; static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1; LEAN_EXPORT uint8_t l_Lean_Parser_hexNumberFn___lambda__1(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_satisfyFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkTokenAndFixPos___closed__2; @@ -190,23 +195,25 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isIdRest___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_pushNone___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notFollowedByFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelseFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_instInhabitedInputContext___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_ParserInfo_collectTokens___default___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkTrailingNode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_atomic___elambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__4; +uint32_t lean_string_utf8_get_fast(lean_object*, lean_object*); lean_object* l_instOrdNat___boxed(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_rawCh___elambda__1(uint32_t, uint8_t, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8; +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16; +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Parser_initCacheForInput___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_tokenFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo(lean_object*); @@ -235,8 +242,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); static lean_object* l_Lean_Parser_checkPrecFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_TokenCacheEntry_startPos___default; LEAN_EXPORT lean_object* l_Lean_Parser_initCacheForInput___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246____boxed(lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkErrorsAt(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9(lean_object*, lean_object*); static lean_object* l_Lean_Parser_instInhabitedParserInfo___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTopFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Parser_TokenMap_insert___spec__2(lean_object*); @@ -271,12 +280,12 @@ LEAN_EXPORT lean_object* l_List_toStringAux___at_Lean_Parser_FirstTokens_toStr__ LEAN_EXPORT lean_object* l_Lean_Parser_andthenInfo___elambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_many1Fn(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_satisfyFn(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_whitespace___closed__4; static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_identFnAux_parse___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_identEqFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkColEq___elambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotExpr___closed__2; @@ -293,6 +302,7 @@ static lean_object* l_Lean_Parser_scientificLitNoAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_indexed(lean_object*); size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_TokenMap_insert___spec__1(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19; lean_object* lean_string_utf8_next(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_invalidLongestMatchParser(lean_object*); static lean_object* l_Lean_Parser_invalidLongestMatchParser___closed__1; @@ -306,13 +316,13 @@ static lean_object* l_Lean_Parser_strLitFnAux___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_octalNumberFn___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_checkColGe___elambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquots(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkColEqFn(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn___lambda__1___closed__2; static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__4; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_eraseReps___at_Lean_Parser_Error_toString___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_takeWhileFn(lean_object*, lean_object*, lean_object*); @@ -323,11 +333,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__2___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_symbolNoAntiquot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___elambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initCacheForInput(lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7(lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); static lean_object* l_Lean_Parser_termParser___closed__1; static lean_object* l_Lean_Parser_fieldIdx___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_insert(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(lean_object*); static lean_object* l_Lean_Parser_longestMatchFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTopFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -349,7 +361,7 @@ static lean_object* l_List_toString___at_Lean_Parser_FirstTokens_toStr___spec__1 LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo___elambda__2(lean_object*); static lean_object* l_Lean_Parser_strLitFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Error_instToStringError; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1(lean_object*); +lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); static lean_object* l_Lean_Parser_strLitFn___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__6(lean_object*); static lean_object* l_Lean_Parser_FirstTokens_instToStringFirstTokens___closed__1; @@ -362,10 +374,9 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser(lean_object*, lean_object static lean_object* l_Lean_Parser_TokenMap_instForInTokenMapProdNameList___closed__2; static lean_object* l_List_toString___at_Lean_Parser_FirstTokens_toStr___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610_(uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_sepByElemParser___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_takeUntilFn___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10; static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkPrec___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_toCtorIdx___boxed(lean_object*); @@ -387,7 +398,7 @@ lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_setExpected(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepLatest(lean_object*, lean_object*); static lean_object* l_Lean_Parser_instInhabitedParserInfo___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser; static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__9; static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBefore(lean_object*); @@ -404,6 +415,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_finishCommentBlock___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBefore___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgs(lean_object*); static lean_object* l_Lean_Parser_whitespace___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_next_x27(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___elambda__1(lean_object*); @@ -416,17 +428,16 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepLatest___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_skip___closed__1; -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_toErrorMsg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLitNoAntiquot___closed__2; LEAN_EXPORT uint8_t l_Lean_Parser_whitespace___lambda__1(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquot(lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9; static lean_object* l_Lean_Parser_fieldIdx___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepNewError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_antiquotExpr; static lean_object* l_Lean_Parser_instInhabitedParserCategory___closed__1; @@ -441,11 +452,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___lambda__1___boxed(lean_obj LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEqFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_optionalFn___closed__2; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withPositionAfterLinebreak(lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numLitFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -453,11 +463,14 @@ lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_mkResult(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instReprLeadingIdentBehavior; LEAN_EXPORT lean_object* l_Lean_Parser_optionaInfo(lean_object*); +lean_object* lean_string_utf8_next_fast(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_FirstTokens_seq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_peekToken(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Error_toString___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_SyntaxNodeKindSet_insert(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_atomic(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__4; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); @@ -468,18 +481,19 @@ static lean_object* l_Lean_Parser_charLitNoAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_peekTokenAux(lean_object*, lean_object*); static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser___boxed(lean_object*, lean_object*); +lean_object* l_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFn___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__3; static lean_object* l_Lean_Parser_TokenMap_instForInTokenMapProdNameList___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkLhsPrecFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_shrinkStack(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_checkTailNoWs(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkTailWs___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchFnAux_parse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_argPrec; LEAN_EXPORT lean_object* l_Lean_Parser_charLitFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoopStep(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4; static lean_object* l_Lean_Parser_sepByElemParser___closed__1; static lean_object* l_Lean_Parser_charLitNoAntiquot___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Parser_TokenMap_insert___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -487,7 +501,6 @@ static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchStep(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instAndThenParser(lean_object*, lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_setExpectedFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_strLitFn(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__12; @@ -505,7 +518,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkWsBefore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBeforeFn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBMap_instForInRBMapProd___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_setCache(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__10; static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_mkIdResult(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -522,22 +535,24 @@ LEAN_EXPORT lean_object* l_Lean_Parser_getNext___boxed(lean_object*, lean_object lean_object* lean_format_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Parser_termParser___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_decQuotDepth___elambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_errorFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_maxPrec; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withPositionAfterLinebreak___elambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19; lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_instBEqOrElseOnAntiquotBehavior___closed__1; static lean_object* l_Lean_Parser_decimalNumberFn___closed__4; static lean_object* l_Lean_Parser_identFn___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepthFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_FirstTokens_merge(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__1; static lean_object* l_Lean_Parser_many1Unbox___closed__1; -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_pushNone___elambda__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_scientificLitFn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1; size_t lean_usize_shift_left(size_t, size_t); lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_leadingParsers___default; @@ -557,7 +572,9 @@ LEAN_EXPORT lean_object* l_Lean_Parser_symbolNoAntiquot___elambda__1(lean_object LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_instInhabitedTokenMap(lean_object*); uint32_t lean_string_utf8_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__6(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__13; +uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPosFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolFn(lean_object*, lean_object*, lean_object*); @@ -566,12 +583,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_hasError___boxed(lean_object* LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nodeInfo___elambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5; static lean_object* l_Lean_Parser_strLitNoAntiquot___closed__3; static lean_object* l_Lean_Parser_manyAux___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLitNoAntiquot___closed__1; +LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_orelseFnCore___closed__1; +static lean_object* l_Lean_Parser_identFnAux_parse___closed__2; static lean_object* l_List_toString___at_Lean_Parser_FirstTokens_toStr___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_lookahead___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserContext_quotDepth___default; @@ -580,8 +598,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_takeUntilFn(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_TokenMap_insert___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo___elambda__1___boxed(lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7; +lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_anyOfFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_incQuotDepth(lean_object*); @@ -592,7 +609,6 @@ static lean_object* l_Lean_Parser_anyOfFn___closed__1; static lean_object* l_Lean_Parser_Error_toString___closed__2; extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l_Lean_Parser_longestMatchStep___closed__2; -static lean_object* l_Lean_Parser_identFnAux_parse___lambda__2___closed__1; static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchMkResult(lean_object*, lean_object*); static lean_object* l_Lean_Parser_orelseFnCore___lambda__1___closed__1; @@ -600,6 +616,7 @@ size_t lean_usize_mul(size_t, size_t); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2___closed__2; static lean_object* l_Lean_Parser_instReprLeadingIdentBehavior___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_whitespace(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_identFnAux_parse___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkLhsPrec___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_rawCh(uint32_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -608,13 +625,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkColEq(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByInfo___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_quotedCharCoreFn___closed__1; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_restore___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_identFnAux_parse___lambda__2___closed__2; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_restore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot___elambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -628,6 +645,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepByFn___boxed(lean_object*, lean_object static lean_object* l_Lean_Parser_FirstTokens_toStr___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_manyAux(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12; uint8_t l_Lean_Syntax_isAntiquot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optionalNoAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__7; @@ -648,7 +666,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__2(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_checkNoImmediateColon; static lean_object* l_Lean_Parser_scientificLitNoAntiquot___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_withResultOfFn(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdentNoAntiquot___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_symbolFnAux___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_whitespace___closed__5; @@ -656,11 +675,10 @@ size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTop___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Int_toNat(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFn(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_instHashableCategoryCacheKey___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEqFn___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_updateCache(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1Unbox___lambda__1___boxed(lean_object*); lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParser___lambda__1(lean_object*, lean_object*); @@ -678,7 +696,7 @@ static lean_object* l_Lean_Parser_scientificLitNoAntiquot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_orelseInfo___elambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_pushNone___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_suppressInsideQuot(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Fn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_noFirstTokenInfo___elambda__1(lean_object*, lean_object*); @@ -693,15 +711,18 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkErrorAt(lean_object*, lean_ static lean_object* l_Lean_Parser_fieldIdxFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos___elambda__1(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbol___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_1115____at_Lean_Parser_mkTokenAndFixPos___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLitFn___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_hexDigitFn(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_checkNoImmediateColon___elambda__1___boxed(lean_object*, lean_object*); @@ -719,17 +740,22 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withoutPosition___elambda__1(lean_object* lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l_Lean_Parser_fieldIdxFn___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepNewError___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__8___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__4___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_trailingTable___default; LEAN_EXPORT lean_object* l_Lean_Parser_ParserModuleContext_currNamespace___default; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_setError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkUnexpectedErrorAt(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withoutForbidden___elambda__1(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Id_instMonadId; LEAN_EXPORT lean_object* l_Lean_Parser_mergeOrElseErrors(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_skip___elambda__1___rarg(lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); @@ -747,6 +773,7 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1NoAntiquot(lean_object*); +static lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_TokenCacheEntry_token___default; static lean_object* l_Lean_Parser_charLitFnAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -754,9 +781,10 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___ LEAN_EXPORT lean_object* l_Lean_Parser_chFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numLitNoAntiquot; static lean_object* l_Lean_Parser_Error_instToStringError___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_instEmptyCollectionTokenMap(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelseInfo___elambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586_(uint8_t, uint8_t); +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__8(lean_object*); uint8_t l_Lean_RBNode_isRed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos(lean_object*, uint8_t); @@ -766,14 +794,18 @@ static lean_object* l_Lean_Parser_mkAntiquot___closed__14; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_FirstTokens_toStr___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withoutInfo___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserModuleContext_openDecls___default; LEAN_EXPORT lean_object* l_Lean_Parser_withResultOf(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notFollowedBy___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_satisfySymbolFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo(lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Parser_categoryParser___elambda__1___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTop(lean_object*, lean_object*); lean_object* l_String_decLt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedPrattParsingTables; @@ -782,6 +814,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Fn(uint8_t, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo; static lean_object* l_Lean_Parser_strLitFnAux___closed__2; +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Parser_categoryParser___elambda__1___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toString___at_Lean_Parser_FirstTokens_toStr___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__4___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_epsilonInfo___closed__3; @@ -798,6 +831,7 @@ static lean_object* l_Lean_Parser_FirstTokens_toStr___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_checkLinebreakBefore___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_rawIdentNoAntiquot; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_replaceLongest(lean_object*, lean_object*); +lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notFollowedBy___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instBEqError; @@ -811,6 +845,7 @@ static lean_object* l_Lean_Parser_unicodeSymbolFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ParserContext_forbiddenTk_x3f___default; LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBefore___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_instHashableCategoryCacheKey; LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_trailingParsers___default; LEAN_EXPORT lean_object* l_Lean_Parser_node___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_noFirstTokenInfo___elambda__2(lean_object*, lean_object*); @@ -819,13 +854,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withoutPosition(lean_object*); static lean_object* l_Lean_Parser_nameLitFn___closed__2; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Parser_Error_toString___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn_parseOptExp___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedFirstTokens; static lean_object* l_Lean_Parser_epsilonInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_rawCh___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6; static lean_object* l_Lean_Parser_numberFnAux___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +static lean_object* l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1; static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_instInhabitedParser___closed__2; @@ -840,13 +874,13 @@ static lean_object* l_Lean_Parser_epsilonInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkPrec___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkPrec(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1(lean_object*); static lean_object* l_Lean_Parser_checkNoImmediateColon___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot(lean_object*, lean_object*, uint8_t); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserFn___rarg(lean_object*); static lean_object* l_Lean_Parser_nameLitFn___closed__4; +lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Parser_info___default___elambda__2(lean_object*); lean_object* l_String_trim(lean_object*); lean_object* l_Lean_isIdEndEscape___boxed(lean_object*); @@ -857,7 +891,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_rawFn(lean_object*, uint8_t, lean_object* static lean_object* l_Lean_Parser_charLitFn___closed__1; lean_object* l_Lean_Syntax_getTailInfo(lean_object*); static lean_object* l_Lean_Parser_instInhabitedParserInfo___closed__2; -static lean_object* l_Lean_Parser_instInhabitedInputContext___closed__4; lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nodeFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -868,16 +901,19 @@ static lean_object* l_Lean_Parser_nameLitFn___closed__1; lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_insert___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTop___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9; +lean_object* lean_mk_array(lean_object*, lean_object*); static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__4; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_toCtorIdx(uint8_t); static lean_object* l_Lean_Parser_instInhabitedInputContext___closed__2; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_strAux_parse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isIdFirst(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_incQuotDepth___elambda__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquot___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__3; static lean_object* l_Lean_Parser_chFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_takeWhile1Fn(lean_object*, lean_object*, lean_object*, lean_object*); @@ -888,29 +924,29 @@ LEAN_EXPORT lean_object* l_Lean_Parser_errorFn___boxed(lean_object*, lean_object static lean_object* l_Lean_Parser_charLitFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_many1Unbox(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByInfo___elambda__2(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_checkColGt(lean_object*); static lean_object* l_Lean_Parser_symbolInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__3(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__6; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_instInhabitedLeadingIdentBehavior; LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_optionaInfo___elambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn_parseOptExp(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_identFnAux_parse___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_leadingNode(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserInfo___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_lexOrd___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__3(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_ParserInfo_collectTokens___default(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_toErrorMsg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__3___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_eoiFn___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAtomicInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkColGtFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__4___rarg(lean_object*, lean_object*, lean_object*); @@ -934,6 +970,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrec(lean_object*); static lean_object* l_Lean_Parser_numLitFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_orelse(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_TokenMap_insert___spec__1___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolFnAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -943,6 +980,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_finishCommentBlock_eoi(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelse___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo___elambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withResultOfInfo___elambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_node(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__8; @@ -953,9 +991,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepPrevError___boxed(lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_strAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instInhabitedParserFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolNoAntiquot(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Parser_categoryParser___elambda__1___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instOrElseParser(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); static lean_object* l_Lean_Parser_instInhabitedParser___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__6___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkTokenAndFixPos(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo___elambda__1___boxed(lean_object*); @@ -972,8 +1012,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkAtom(lean_object*, lean_object*); static lean_object* l_Lean_Parser_pushNone___elambda__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mkEOIError(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Parser_initCacheForInput___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_updateTokenCache(lean_object*, lean_object*); lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_indexed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_rawFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserOfStackFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLitFn___closed__3; @@ -987,6 +1030,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_TokenCacheEntry_stopPos___default; static lean_object* l_Lean_Parser_indexed___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_unicodeSymbolInfo___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_tokenFnAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot___elambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_tokenWithAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); @@ -1007,6 +1051,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__6__ LEAN_EXPORT lean_object* l_Lean_Parser_identFn(lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_1115____at_Lean_Parser_ParserState_hasError___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9___boxed(lean_object*, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1016,6 +1061,7 @@ static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNo LEAN_EXPORT lean_object* l_Lean_Parser_charLitNoAntiquot; lean_object* l_Lean_mkErrorStringWithPos(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_fieldIdxFn___closed__2; +extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolFnAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); @@ -1024,18 +1070,17 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkColGtFn___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Parser_checkColEq___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Info___elambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1; static lean_object* l_Lean_Parser_mkAtomicInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_suppressInsideQuot___elambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_identNoAntiquot___closed__2; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_withoutForbidden(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_rawAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mergeErrors(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_fieldIdxFn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10; static lean_object* l_Lean_Parser_hexNumberFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_mkIdent(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -1043,7 +1088,6 @@ lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_longestMatchStep___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_error(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepTop___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_toStringAux___at_Lean_Parser_FirstTokens_toStr___spec__2(uint8_t, lean_object*); @@ -1155,6 +1199,214 @@ x_1 = lean_box(0); return x_1; } } +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_1, 2); +x_6 = lean_ctor_get(x_2, 0); +x_7 = lean_ctor_get(x_2, 1); +x_8 = lean_ctor_get(x_2, 2); +x_9 = lean_name_eq(x_3, x_6); +if (x_9 == 0) +{ +uint8_t x_10; +x_10 = 0; +return x_10; +} +else +{ +uint8_t x_11; +x_11 = lean_nat_dec_eq(x_4, x_7); +if (x_11 == 0) +{ +uint8_t x_12; +x_12 = 0; +return x_12; +} +else +{ +uint8_t x_13; +x_13 = lean_nat_dec_eq(x_5, x_8); +return x_13; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_instBEqCategoryCacheKey___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_instBEqCategoryCacheKey() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_instBEqCategoryCacheKey___closed__1; +return x_1; +} +} +LEAN_EXPORT uint64_t l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; uint64_t x_5; uint64_t x_6; uint64_t x_7; uint64_t x_8; uint64_t x_9; uint64_t x_10; uint64_t x_11; +x_2 = lean_ctor_get(x_1, 0); +x_3 = lean_ctor_get(x_1, 1); +x_4 = lean_ctor_get(x_1, 2); +x_5 = 0; +x_6 = l_Lean_Name_hash___override(x_2); +x_7 = lean_uint64_mix_hash(x_5, x_6); +x_8 = lean_uint64_of_nat(x_3); +x_9 = lean_uint64_mix_hash(x_7, x_8); +x_10 = lean_uint64_of_nat(x_4); +x_11 = lean_uint64_mix_hash(x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246____boxed(lean_object* x_1) { +_start: +{ +uint64_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_1); +lean_dec(x_1); +x_3 = lean_box_uint64(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_instHashableCategoryCacheKey___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246____boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_instHashableCategoryCacheKey() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_instHashableCategoryCacheKey___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Error_unexpected___default___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Error_unexpected___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Error_unexpected___default___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Error_expected___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_instInhabitedError___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Error_unexpected___default___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_instInhabitedError() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_instInhabitedError___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_1, 1); +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_string_dec_eq(x_3, x_5); +if (x_7 == 0) +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +else +{ +uint8_t x_9; +x_9 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3154____spec__1(x_4, x_6); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_instBEqError___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_instBEqError() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_instBEqError___closed__1; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Parser_initCacheForInput___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_mkHashMapImp___rarg(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_Parser_initCacheForInput___closed__1() { _start: { @@ -1167,7 +1419,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_initCacheForInput(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_2 = lean_string_utf8_byte_size(x_1); x_3 = l_Lean_Parser_initCacheForInput___closed__1; x_4 = lean_nat_add(x_2, x_3); @@ -1178,7 +1430,21 @@ x_7 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_7, 0, x_4); lean_ctor_set(x_7, 1, x_5); lean_ctor_set(x_7, 2, x_6); -return x_7; +x_8 = lean_unsigned_to_nat(8u); +x_9 = l_Lean_mkHashMapImp___rarg(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Parser_initCacheForInput___spec__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_mkHashMap___at_Lean_Parser_initCacheForInput___spec__1(x_1); +lean_dec(x_1); +return x_2; } } LEAN_EXPORT lean_object* l_Lean_Parser_initCacheForInput___boxed(lean_object* x_1) { @@ -1797,16 +2063,8 @@ return x_2; static lean_object* _init_l_Lean_Parser_instInhabitedInputContext___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_instInhabitedInputContext___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_1 = l_Lean_Parser_Error_unexpected___default___closed__1; x_2 = l_Lean_Parser_instInhabitedInputContext___closed__1; x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -1815,12 +2073,12 @@ lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_instInhabitedInputContext___closed__4() { +static lean_object* _init_l_Lean_Parser_instInhabitedInputContext___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_instInhabitedInputContext___closed__2; -x_2 = l_Lean_Parser_instInhabitedInputContext___closed__3; +x_1 = l_Lean_Parser_Error_unexpected___default___closed__1; +x_2 = l_Lean_Parser_instInhabitedInputContext___closed__2; x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_1); @@ -1832,7 +2090,7 @@ static lean_object* _init_l_Lean_Parser_instInhabitedInputContext() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_instInhabitedInputContext___closed__4; +x_1 = l_Lean_Parser_instInhabitedInputContext___closed__3; return x_1; } } @@ -1884,131 +2142,45 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Error_unexpected___default() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_instInhabitedInputContext___closed__2; -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Error_expected___default() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__1() { _start: { lean_object* x_1; -x_1 = lean_box(0); +x_1 = lean_mk_string_from_bytes(" or ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_instInhabitedError___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_instInhabitedInputContext___closed__2; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_instInhabitedError() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__2() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_instInhabitedError___closed__1; +x_1 = lean_mk_string_from_bytes(", ", 2); return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString(lean_object* x_1) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_ctor_get(x_1, 1); -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -x_7 = lean_string_dec_eq(x_3, x_5); -if (x_7 == 0) +if (lean_obj_tag(x_1) == 0) { -uint8_t x_8; -x_8 = 0; -return x_8; +lean_object* x_2; +x_2 = l_Lean_Parser_Error_unexpected___default___closed__1; +return x_2; } else { -uint8_t x_9; -x_9 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3154____spec__1(x_4, x_6); -return x_9; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357____boxed(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 0) { -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(x_1, x_2); -lean_dec(x_2); +lean_object* x_4; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); lean_dec(x_1); -x_4 = lean_box(x_3); return x_4; } -} -static lean_object* _init_l_Lean_Parser_instBEqError___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357____boxed), 2, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_instBEqError() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_instBEqError___closed__1; -return x_1; -} -} -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" or ", 4); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(", ", 2); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_2; -x_2 = l_Lean_Parser_instInhabitedInputContext___closed__2; -return x_2; -} -else -{ -lean_object* x_3; -x_3 = lean_ctor_get(x_1, 1); -lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -return x_4; -} -else +else { lean_object* x_5; x_5 = lean_ctor_get(x_3, 1); @@ -2346,7 +2518,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_1, 1); lean_inc(x_3); lean_dec(x_1); -x_4 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_4 = l_Lean_Parser_Error_unexpected___default___closed__1; x_5 = lean_string_dec_eq(x_2, x_4); x_6 = lean_box(0); x_7 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3154____spec__1(x_3, x_6); @@ -2471,7 +2643,7 @@ lean_inc(x_3); x_4 = lean_ctor_get(x_2, 1); lean_inc(x_4); lean_dec(x_2); -x_5 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_5 = l_Lean_Parser_Error_unexpected___default___closed__1; x_6 = lean_string_dec_eq(x_3, x_5); x_7 = !lean_is_exclusive(x_1); if (x_7 == 0) @@ -2587,7 +2759,7 @@ else lean_object* x_6; lean_object* x_7; uint8_t x_8; x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_2, 0); -x_8 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(x_6, x_7); +x_8 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(x_6, x_7); return x_8; } } @@ -2944,6 +3116,53 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_next_x27(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_1, 2); +lean_dec(x_6); +x_7 = lean_string_utf8_next_fast(x_2, x_3); +lean_ctor_set(x_1, 2, x_7); +return x_1; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_8 = lean_ctor_get(x_1, 0); +x_9 = lean_ctor_get(x_1, 1); +x_10 = lean_ctor_get(x_1, 3); +x_11 = lean_ctor_get(x_1, 4); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_1); +x_12 = lean_string_utf8_next_fast(x_2, x_3); +x_13 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_13, 0, x_8); +lean_ctor_set(x_13, 1, x_9); +lean_ctor_set(x_13, 2, x_12); +lean_ctor_set(x_13, 3, x_10); +lean_ctor_set(x_13, 4, x_11); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_next_x27___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Parser_ParserState_next_x27(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_toErrorMsg(lean_object* x_1, lean_object* x_2) { _start: { @@ -2954,7 +3173,7 @@ if (lean_obj_tag(x_3) == 0) { lean_object* x_4; lean_dec(x_2); -x_4 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_4 = l_Lean_Parser_Error_unexpected___default___closed__1; return x_4; } else @@ -3226,7 +3445,7 @@ x_5 = lean_box(0); x_6 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_6, 0, x_2); lean_ctor_set(x_6, 1, x_5); -x_7 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_7 = l_Lean_Parser_Error_unexpected___default___closed__1; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); @@ -3251,7 +3470,7 @@ x_14 = lean_box(0); x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_2); lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_16 = l_Lean_Parser_Error_unexpected___default___closed__1; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -3284,7 +3503,7 @@ x_8 = lean_box(0); x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_2); lean_ctor_set(x_9, 1, x_8); -x_10 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_10 = l_Lean_Parser_Error_unexpected___default___closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -3312,7 +3531,7 @@ x_19 = lean_box(0); x_20 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_20, 0, x_2); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_21 = l_Lean_Parser_Error_unexpected___default___closed__1; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -3417,7 +3636,7 @@ x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_2); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_13 = l_Lean_Parser_Error_unexpected___default___closed__1; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -3444,7 +3663,7 @@ x_21 = lean_box(0); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_2); lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_23 = l_Lean_Parser_Error_unexpected___default___closed__1; x_24 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); @@ -3484,7 +3703,7 @@ x_36 = lean_box(0); x_37 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_37, 0, x_2); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_38 = l_Lean_Parser_Error_unexpected___default___closed__1; x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_37); @@ -3508,7 +3727,7 @@ x_44 = lean_box(0); x_45 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_45, 0, x_2); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_46 = l_Lean_Parser_Error_unexpected___default___closed__1; x_47 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_47, 0, x_46); lean_ctor_set(x_47, 1, x_45); @@ -3547,7 +3766,7 @@ x_57 = lean_box(0); x_58 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_58, 0, x_2); lean_ctor_set(x_58, 1, x_57); -x_59 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_59 = l_Lean_Parser_Error_unexpected___default___closed__1; x_60 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_60, 0, x_59); lean_ctor_set(x_60, 1, x_58); @@ -3585,7 +3804,7 @@ x_8 = lean_ctor_get(x_1, 2); lean_dec(x_8); x_9 = lean_box(0); x_10 = lean_array_push(x_6, x_9); -x_11 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_11 = l_Lean_Parser_Error_unexpected___default___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_2); @@ -3608,7 +3827,7 @@ lean_inc(x_14); lean_dec(x_1); x_17 = lean_box(0); x_18 = lean_array_push(x_14, x_17); -x_19 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_19 = l_Lean_Parser_Error_unexpected___default___closed__1; x_20 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_2); @@ -3644,7 +3863,7 @@ x_29 = l_Array_shrink___rarg(x_24, x_28); lean_dec(x_28); x_30 = lean_box(0); x_31 = lean_array_push(x_29, x_30); -x_32 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_32 = l_Lean_Parser_Error_unexpected___default___closed__1; x_33 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_33, 0, x_32); lean_ctor_set(x_33, 1, x_2); @@ -3664,7 +3883,7 @@ x_35 = l_Array_shrink___rarg(x_24, x_34); lean_dec(x_34); x_36 = lean_box(0); x_37 = lean_array_push(x_35, x_36); -x_38 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_38 = l_Lean_Parser_Error_unexpected___default___closed__1; x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_2); @@ -3699,7 +3918,7 @@ x_46 = l_Array_shrink___rarg(x_41, x_44); lean_dec(x_44); x_47 = lean_box(0); x_48 = lean_array_push(x_46, x_47); -x_49 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_49 = l_Lean_Parser_Error_unexpected___default___closed__1; x_50 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_2); @@ -4106,7 +4325,7 @@ if (x_1 == 0) if (lean_obj_tag(x_2) == 0) { lean_object* x_3; -x_3 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_3 = l_Lean_Parser_Error_unexpected___default___closed__1; return x_3; } else @@ -4132,7 +4351,7 @@ else if (lean_obj_tag(x_2) == 0) { lean_object* x_11; -x_11 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_11 = l_Lean_Parser_Error_unexpected___default___closed__1; return x_11; } else @@ -4579,7 +4798,7 @@ if (x_1 == 0) if (lean_obj_tag(x_2) == 0) { lean_object* x_3; -x_3 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_3 = l_Lean_Parser_Error_unexpected___default___closed__1; return x_3; } else @@ -4610,7 +4829,7 @@ else if (lean_obj_tag(x_2) == 0) { lean_object* x_16; -x_16 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_16 = l_Lean_Parser_Error_unexpected___default___closed__1; return x_16; } else @@ -4755,7 +4974,7 @@ lean_inc(x_5); x_6 = lean_array_get_size(x_5); lean_dec(x_5); x_7 = lean_apply_2(x_2, x_3, x_4); -x_8 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_8 = l_Lean_Parser_Error_unexpected___default___closed__1; x_9 = lean_string_append(x_8, x_1); x_10 = l_Lean_Parser_dbgTraceStateFn___closed__1; x_11 = lean_string_append(x_9, x_10); @@ -6403,7 +6622,7 @@ x_6 = l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -6415,7 +6634,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -6423,7 +6642,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942_(x_3, x_4); +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -6432,7 +6651,7 @@ static lean_object* _init_l_Lean_Parser_instBEqOrElseOnAntiquotBehavior___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221____boxed), 2, 0); return x_1; } } @@ -6510,7 +6729,7 @@ x_14 = l_Lean_instInhabitedSyntax; x_15 = l_Array_back___rarg(x_14, x_10); x_16 = l_Lean_Parser_orelseFnCore___closed__1; x_17 = 0; -x_18 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942_(x_3, x_17); +x_18 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221_(x_3, x_17); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; @@ -6568,7 +6787,7 @@ if (x_27 == 0) { uint8_t x_28; uint8_t x_29; x_28 = 2; -x_29 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_2942_(x_3, x_28); +x_29 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_3221_(x_3, x_28); if (x_29 == 0) { lean_dec(x_24); @@ -7259,7 +7478,7 @@ x_9 = l_Lean_Parser_ParserState_restore(x_7, x_5, x_6); lean_dec(x_5); x_10 = l_Lean_Parser_notFollowedByFn___closed__1; x_11 = lean_string_append(x_10, x_2); -x_12 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_12 = l_Lean_Parser_Error_unexpected___default___closed__1; x_13 = lean_string_append(x_11, x_12); x_14 = lean_box(0); x_15 = l_Lean_Parser_ParserState_mkUnexpectedError(x_9, x_13, x_14); @@ -8205,7 +8424,7 @@ x_8 = lean_string_utf8_at_end(x_7, x_5); if (x_8 == 0) { uint32_t x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_string_utf8_get(x_7, x_5); +x_9 = lean_string_utf8_get_fast(x_7, x_5); x_10 = lean_box_uint32(x_9); x_11 = lean_apply_1(x_1, x_10); x_12 = lean_unbox(x_11); @@ -8222,7 +8441,7 @@ else { lean_object* x_15; lean_dec(x_2); -x_15 = l_Lean_Parser_ParserState_next(x_4, x_7, x_5); +x_15 = l_Lean_Parser_ParserState_next_x27(x_4, x_7, x_5, lean_box(0)); lean_dec(x_5); return x_15; } @@ -8261,7 +8480,7 @@ x_7 = lean_string_utf8_at_end(x_6, x_4); if (x_7 == 0) { uint32_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_string_utf8_get(x_6, x_4); +x_8 = lean_string_utf8_get_fast(x_6, x_4); x_9 = lean_box_uint32(x_8); lean_inc(x_1); x_10 = lean_apply_1(x_1, x_9); @@ -8270,7 +8489,7 @@ lean_dec(x_10); if (x_11 == 0) { lean_object* x_12; -x_12 = l_Lean_Parser_ParserState_next(x_3, x_6, x_4); +x_12 = l_Lean_Parser_ParserState_next_x27(x_3, x_6, x_4, lean_box(0)); lean_dec(x_4); x_3 = x_12; goto _start; @@ -8395,8 +8614,8 @@ x_7 = lean_string_utf8_at_end(x_5, x_6); if (x_7 == 0) { uint32_t x_8; lean_object* x_9; uint32_t x_10; uint8_t x_11; -x_8 = lean_string_utf8_get(x_5, x_6); -x_9 = lean_string_utf8_next(x_5, x_6); +x_8 = lean_string_utf8_get_fast(x_5, x_6); +x_9 = lean_string_utf8_next_fast(x_5, x_6); lean_dec(x_6); x_10 = 45; x_11 = lean_uint32_dec_eq(x_8, x_10); @@ -8419,7 +8638,7 @@ x_16 = lean_string_utf8_at_end(x_5, x_9); if (x_16 == 0) { uint32_t x_17; uint8_t x_18; -x_17 = lean_string_utf8_get(x_5, x_9); +x_17 = lean_string_utf8_get_fast(x_5, x_9); x_18 = lean_uint32_dec_eq(x_17, x_10); if (x_18 == 0) { @@ -8434,7 +8653,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; x_21 = lean_unsigned_to_nat(1u); x_22 = lean_nat_add(x_1, x_21); lean_dec(x_1); -x_23 = l_Lean_Parser_ParserState_next(x_3, x_5, x_9); +x_23 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_9, lean_box(0)); lean_dec(x_9); x_1 = x_22; x_3 = x_23; @@ -8458,7 +8677,7 @@ x_26 = lean_string_utf8_at_end(x_5, x_9); if (x_26 == 0) { uint32_t x_27; uint32_t x_28; uint8_t x_29; -x_27 = lean_string_utf8_get(x_5, x_9); +x_27 = lean_string_utf8_get_fast(x_5, x_9); x_28 = 47; x_29 = lean_uint32_dec_eq(x_27, x_28); if (x_29 == 0) @@ -8479,7 +8698,7 @@ if (x_33 == 0) lean_object* x_34; lean_object* x_35; x_34 = lean_nat_sub(x_1, x_32); lean_dec(x_1); -x_35 = l_Lean_Parser_ParserState_next(x_3, x_5, x_9); +x_35 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_9, lean_box(0)); lean_dec(x_9); x_1 = x_34; x_3 = x_35; @@ -8489,7 +8708,7 @@ else { lean_object* x_37; lean_dec(x_1); -x_37 = l_Lean_Parser_ParserState_next(x_3, x_5, x_9); +x_37 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_9, lean_box(0)); lean_dec(x_9); return x_37; } @@ -8592,7 +8811,7 @@ x_6 = lean_string_utf8_at_end(x_4, x_5); if (x_6 == 0) { uint32_t x_7; uint32_t x_8; uint8_t x_9; -x_7 = lean_string_utf8_get(x_4, x_5); +x_7 = lean_string_utf8_get_fast(x_4, x_5); x_8 = 9; x_9 = lean_uint32_dec_eq(x_7, x_8); if (x_9 == 0) @@ -8619,7 +8838,7 @@ return x_2; else { lean_object* x_15; uint32_t x_16; uint8_t x_17; -x_15 = lean_string_utf8_next(x_4, x_5); +x_15 = lean_string_utf8_next_fast(x_4, x_5); lean_dec(x_5); x_16 = lean_string_utf8_get(x_4, x_15); x_17 = lean_uint32_dec_eq(x_16, x_11); @@ -8674,7 +8893,7 @@ return x_2; else { lean_object* x_27; uint32_t x_28; uint8_t x_29; -x_27 = lean_string_utf8_next(x_4, x_5); +x_27 = lean_string_utf8_next_fast(x_4, x_5); lean_dec(x_5); x_28 = lean_string_utf8_get(x_4, x_27); x_29 = lean_uint32_dec_eq(x_28, x_11); @@ -8701,7 +8920,7 @@ return x_33; else { lean_object* x_34; -x_34 = l_Lean_Parser_ParserState_next(x_2, x_4, x_5); +x_34 = l_Lean_Parser_ParserState_next_x27(x_2, x_4, x_5, lean_box(0)); lean_dec(x_5); lean_dec(x_4); x_2 = x_34; @@ -8877,7 +9096,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_obj x_5 = lean_box_uint32(x_1); x_6 = lean_alloc_closure((void*)(l_UInt32_decEq___boxed), 2, 1); lean_closure_set(x_6, 0, x_5); -x_7 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_7 = l_Lean_Parser_Error_unexpected___default___closed__1; x_8 = lean_string_push(x_7, x_1); x_9 = l_Lean_Parser_chFn___closed__1; x_10 = lean_string_append(x_9, x_8); @@ -8970,8 +9189,8 @@ x_6 = lean_string_utf8_at_end(x_4, x_5); if (x_6 == 0) { uint32_t x_7; lean_object* x_8; lean_object* x_9; uint8_t x_22; -x_7 = lean_string_utf8_get(x_4, x_5); -x_8 = lean_string_utf8_next(x_4, x_5); +x_7 = lean_string_utf8_get_fast(x_4, x_5); +x_8 = lean_string_utf8_next_fast(x_4, x_5); lean_dec(x_5); x_22 = l_Char_isDigit(x_7); if (x_22 == 0) @@ -9124,7 +9343,7 @@ x_7 = lean_string_utf8_at_end(x_5, x_6); if (x_7 == 0) { uint32_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_string_utf8_get(x_5, x_6); +x_8 = lean_string_utf8_get_fast(x_5, x_6); x_9 = lean_box_uint32(x_8); x_10 = lean_apply_1(x_1, x_9); x_11 = lean_unbox(x_10); @@ -9153,7 +9372,7 @@ return x_18; else { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = l_Lean_Parser_ParserState_next(x_3, x_5, x_6); +x_19 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_6, lean_box(0)); lean_dec(x_6); lean_dec(x_5); x_20 = l_Lean_Parser_quotedCharCoreFn___closed__2; @@ -9165,7 +9384,7 @@ return x_22; else { lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = l_Lean_Parser_ParserState_next(x_3, x_5, x_6); +x_23 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_6, lean_box(0)); lean_dec(x_6); lean_dec(x_5); x_24 = l_Lean_Parser_quotedCharCoreFn___closed__2; @@ -9177,7 +9396,7 @@ else { lean_object* x_26; lean_dec(x_2); -x_26 = l_Lean_Parser_ParserState_next(x_3, x_5, x_6); +x_26 = l_Lean_Parser_ParserState_next_x27(x_3, x_5, x_6, lean_box(0)); lean_dec(x_6); lean_dec(x_5); return x_26; @@ -9371,8 +9590,8 @@ x_7 = lean_string_utf8_at_end(x_5, x_6); if (x_7 == 0) { uint32_t x_8; lean_object* x_9; lean_object* x_10; uint32_t x_11; uint8_t x_12; -x_8 = lean_string_utf8_get(x_5, x_6); -x_9 = lean_string_utf8_next(x_5, x_6); +x_8 = lean_string_utf8_get_fast(x_5, x_6); +x_9 = lean_string_utf8_next_fast(x_5, x_6); lean_dec(x_6); lean_inc(x_9); x_10 = l_Lean_Parser_ParserState_setPos(x_3, x_9); @@ -9529,8 +9748,8 @@ x_7 = lean_string_utf8_at_end(x_5, x_6); if (x_7 == 0) { uint32_t x_8; lean_object* x_9; lean_object* x_10; uint32_t x_11; uint8_t x_12; -x_8 = lean_string_utf8_get(x_5, x_6); -x_9 = lean_string_utf8_next(x_5, x_6); +x_8 = lean_string_utf8_get_fast(x_5, x_6); +x_9 = lean_string_utf8_next_fast(x_5, x_6); lean_dec(x_6); lean_dec(x_5); x_10 = l_Lean_Parser_ParserState_setPos(x_3, x_9); @@ -10154,7 +10373,7 @@ x_6 = lean_string_utf8_at_end(x_4, x_5); if (x_6 == 0) { uint32_t x_7; uint32_t x_8; uint8_t x_9; -x_7 = lean_string_utf8_get(x_4, x_5); +x_7 = lean_string_utf8_get_fast(x_4, x_5); x_8 = 48; x_9 = lean_uint32_dec_eq(x_7, x_8); if (x_9 == 0) @@ -10183,7 +10402,7 @@ return x_14; else { lean_object* x_15; uint32_t x_16; uint32_t x_17; uint8_t x_18; -x_15 = lean_string_utf8_next(x_4, x_5); +x_15 = lean_string_utf8_next_fast(x_4, x_5); x_16 = lean_string_utf8_get(x_4, x_15); x_17 = 98; x_18 = lean_uint32_dec_eq(x_16, x_17); @@ -10654,39 +10873,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_1, 2); -lean_inc(x_9); -x_10 = l_Lean_Parser_ParserState_next(x_1, x_2, x_9); -x_11 = lean_string_utf8_extract(x_2, x_3, x_9); -lean_dec(x_3); -x_12 = l_Lean_Name_str___override(x_4, x_11); -x_13 = l_Lean_Parser_isIdCont(x_2, x_10); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_9); -lean_dec(x_2); -x_14 = l_Lean_Parser_mkIdResult(x_5, x_6, x_12, x_7, x_10); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_string_utf8_next(x_2, x_9); -lean_dec(x_9); -x_16 = l_Lean_Parser_ParserState_next(x_10, x_2, x_15); -lean_dec(x_15); -lean_dec(x_2); -x_17 = l_Lean_Parser_identFnAux_parse(x_5, x_6, x_12, x_7, x_16); -return x_17; -} -} -} -static lean_object* _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Parser_identFnAux_parse___closed__1() { _start: { lean_object* x_1; @@ -10694,7 +10881,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_isIdRest___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Parser_identFnAux_parse___closed__2() { _start: { lean_object* x_1; @@ -10702,7 +10889,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_isIdEndEscape___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Parser_identFnAux_parse___closed__3() { _start: { lean_object* x_1; @@ -10710,135 +10897,139 @@ x_1 = lean_mk_string_from_bytes("unterminated identifier escape", 30); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint32_t x_9; uint32_t x_10; uint8_t x_11; -lean_dec(x_8); -x_9 = lean_string_utf8_get(x_1, x_2); -x_10 = l_Lean_idBeginEscape; -x_11 = lean_uint32_dec_eq(x_9, x_10); -if (x_11 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +x_9 = lean_string_utf8_at_end(x_7, x_8); +if (x_9 == 0) { -uint8_t x_12; -x_12 = l_Lean_isIdFirst(x_9); +uint32_t x_10; uint32_t x_11; uint8_t x_12; +x_10 = lean_string_utf8_get_fast(x_7, x_8); +x_11 = l_Lean_idBeginEscape; +x_12 = lean_uint32_dec_eq(x_10, x_11); if (x_12 == 0) { -lean_object* x_13; +uint8_t x_13; +x_13 = l_Lean_isIdFirst(x_10); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_8); lean_dec(x_7); -lean_dec(x_2); -lean_dec(x_1); -x_13 = l_Lean_Parser_mkTokenAndFixPos(x_3, x_4, x_5, x_6); -return x_13; +lean_dec(x_3); +x_14 = l_Lean_Parser_mkTokenAndFixPos(x_1, x_2, x_4, x_5); +return x_14; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_14 = l_Lean_Parser_ParserState_next(x_6, x_1, x_2); -x_15 = l_Lean_Parser_identFnAux_parse___lambda__2___closed__1; -x_16 = l_Lean_Parser_takeWhileFn(x_15, x_5, x_14); -x_17 = lean_ctor_get(x_16, 2); -lean_inc(x_17); -x_18 = lean_string_utf8_extract(x_1, x_2, x_17); -lean_dec(x_2); -x_19 = l_Lean_Name_str___override(x_7, x_18); -x_20 = l_Lean_Parser_isIdCont(x_1, x_16); -if (x_20 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_15 = l_Lean_Parser_ParserState_next(x_5, x_7, x_8); +x_16 = l_Lean_Parser_identFnAux_parse___closed__1; +x_17 = l_Lean_Parser_takeWhileFn(x_16, x_4, x_15); +x_18 = lean_ctor_get(x_17, 2); +lean_inc(x_18); +x_19 = lean_string_utf8_extract(x_7, x_8, x_18); +lean_dec(x_8); +x_20 = l_Lean_Name_str___override(x_3, x_19); +x_21 = l_Lean_Parser_isIdCont(x_7, x_17); +if (x_21 == 0) { -lean_object* x_21; -lean_dec(x_17); -lean_dec(x_1); -x_21 = l_Lean_Parser_mkIdResult(x_3, x_4, x_19, x_5, x_16); -return x_21; +lean_object* x_22; +lean_dec(x_18); +lean_dec(x_7); +x_22 = l_Lean_Parser_mkIdResult(x_1, x_2, x_20, x_4, x_17); +return x_22; } else { -lean_object* x_22; lean_object* x_23; -x_22 = l_Lean_Parser_ParserState_next(x_16, x_1, x_17); -lean_dec(x_17); -lean_dec(x_1); -x_23 = l_Lean_Parser_identFnAux_parse(x_3, x_4, x_19, x_5, x_22); -return x_23; +lean_object* x_23; +x_23 = l_Lean_Parser_ParserState_next(x_17, x_7, x_18); +lean_dec(x_18); +lean_dec(x_7); +x_3 = x_20; +x_5 = x_23; +goto _start; } } } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_24 = lean_string_utf8_next(x_1, x_2); -lean_dec(x_2); -lean_inc(x_24); -x_25 = l_Lean_Parser_ParserState_setPos(x_6, x_24); -x_26 = l_Lean_Parser_identFnAux_parse___lambda__2___closed__2; -x_27 = l_Lean_Parser_takeUntilFn(x_26, x_5, x_25); -x_28 = lean_ctor_get(x_27, 2); -lean_inc(x_28); -x_29 = lean_string_utf8_at_end(x_1, x_28); -lean_dec(x_28); -if (x_29 == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_25 = lean_string_utf8_next_fast(x_7, x_8); +lean_dec(x_8); +lean_inc(x_25); +x_26 = l_Lean_Parser_ParserState_setPos(x_5, x_25); +x_27 = l_Lean_Parser_identFnAux_parse___closed__2; +x_28 = l_Lean_Parser_takeUntilFn(x_27, x_4, x_26); +x_29 = lean_ctor_get(x_28, 2); +lean_inc(x_29); +x_30 = lean_string_utf8_at_end(x_7, x_29); +if (x_30 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = lean_box(0); -x_31 = l_Lean_Parser_identFnAux_parse___lambda__1(x_27, x_1, x_24, x_7, x_3, x_4, x_5, x_30); -return x_31; +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = l_Lean_Parser_ParserState_next_x27(x_28, x_7, x_29, lean_box(0)); +x_32 = lean_string_utf8_extract(x_7, x_25, x_29); +lean_dec(x_25); +x_33 = l_Lean_Name_str___override(x_3, x_32); +x_34 = l_Lean_Parser_isIdCont(x_7, x_31); +if (x_34 == 0) +{ +lean_object* x_35; +lean_dec(x_29); +lean_dec(x_7); +x_35 = l_Lean_Parser_mkIdResult(x_1, x_2, x_33, x_4, x_31); +return x_35; } else { -lean_object* x_32; lean_object* x_33; +lean_object* x_36; lean_object* x_37; +x_36 = lean_string_utf8_next_fast(x_7, x_29); +lean_dec(x_29); +x_37 = l_Lean_Parser_ParserState_next(x_31, x_7, x_36); +lean_dec(x_36); +lean_dec(x_7); +x_3 = x_33; +x_5 = x_37; +goto _start; +} +} +else +{ +lean_object* x_39; lean_object* x_40; +lean_dec(x_29); lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -x_32 = l_Lean_Parser_identFnAux_parse___lambda__2___closed__3; -x_33 = l_Lean_Parser_ParserState_mkUnexpectedErrorAt(x_27, x_32, x_24); -return x_33; -} -} +x_39 = l_Lean_Parser_identFnAux_parse___closed__3; +x_40 = l_Lean_Parser_ParserState_mkUnexpectedErrorAt(x_28, x_39, x_25); +return x_40; } } -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_ctor_get(x_5, 2); -lean_inc(x_8); -x_9 = lean_string_utf8_at_end(x_7, x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_box(0); -x_11 = l_Lean_Parser_identFnAux_parse___lambda__2(x_7, x_8, x_1, x_2, x_4, x_5, x_3, x_10); -return x_11; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_12 = lean_box(0); -x_13 = l_Lean_Parser_ParserState_mkEOIError___closed__1; -x_14 = l_Lean_Parser_ParserState_mkUnexpectedError(x_5, x_13, x_12); -return x_14; -} -} +x_41 = lean_box(0); +x_42 = l_Lean_Parser_ParserState_mkEOIError___closed__1; +x_43 = l_Lean_Parser_ParserState_mkUnexpectedError(x_5, x_42, x_41); +return x_43; } -LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Parser_identFnAux_parse___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_8); -return x_9; } } LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -11050,86 +11241,157 @@ return x_31; } } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_updateCache(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_updateTokenCache(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; -x_3 = lean_ctor_get(x_2, 4); +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_2, 3); lean_inc(x_3); -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_4 = lean_ctor_get(x_2, 0); +x_4 = lean_ctor_get(x_2, 4); lean_inc(x_4); -x_5 = lean_ctor_get(x_2, 1); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_2, 0); lean_inc(x_5); -x_6 = lean_ctor_get(x_2, 2); +x_6 = lean_ctor_get(x_2, 1); lean_inc(x_6); -x_7 = lean_array_get_size(x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_eq(x_7, x_8); -lean_dec(x_7); -if (x_9 == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_2); -if (x_10 == 0) +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_7); +x_8 = !lean_is_exclusive(x_3); +if (x_8 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_11 = lean_ctor_get(x_2, 4); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_3, 1); +x_10 = lean_ctor_get(x_3, 0); +lean_dec(x_10); +x_11 = lean_array_get_size(x_5); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_nat_dec_eq(x_11, x_12); lean_dec(x_11); -x_12 = lean_ctor_get(x_2, 3); -lean_dec(x_12); -x_13 = lean_ctor_get(x_2, 2); -lean_dec(x_13); -x_14 = lean_ctor_get(x_2, 1); -lean_dec(x_14); -x_15 = lean_ctor_get(x_2, 0); -lean_dec(x_15); -x_16 = l_Lean_instInhabitedSyntax; -x_17 = l_Array_back___rarg(x_16, x_4); -lean_inc(x_6); -x_18 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_6); -lean_ctor_set(x_18, 2, x_17); -x_19 = lean_box(0); -lean_ctor_set(x_2, 4, x_19); -lean_ctor_set(x_2, 3, x_18); -return x_2; -} -else +if (x_13 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_dec(x_2); +uint8_t x_14; +x_14 = !lean_is_exclusive(x_2); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_ctor_get(x_2, 4); +lean_dec(x_15); +x_16 = lean_ctor_get(x_2, 3); +lean_dec(x_16); +x_17 = lean_ctor_get(x_2, 2); +lean_dec(x_17); +x_18 = lean_ctor_get(x_2, 1); +lean_dec(x_18); +x_19 = lean_ctor_get(x_2, 0); +lean_dec(x_19); x_20 = l_Lean_instInhabitedSyntax; -x_21 = l_Array_back___rarg(x_20, x_4); -lean_inc(x_6); +x_21 = l_Array_back___rarg(x_20, x_5); +lean_inc(x_7); x_22 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_22, 0, x_1); -lean_ctor_set(x_22, 1, x_6); +lean_ctor_set(x_22, 1, x_7); lean_ctor_set(x_22, 2, x_21); +lean_ctor_set(x_3, 0, x_22); x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_24, 0, x_4); -lean_ctor_set(x_24, 1, x_5); -lean_ctor_set(x_24, 2, x_6); -lean_ctor_set(x_24, 3, x_22); -lean_ctor_set(x_24, 4, x_23); -return x_24; +lean_ctor_set(x_2, 4, x_23); +return x_2; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_dec(x_2); +x_24 = l_Lean_instInhabitedSyntax; +x_25 = l_Array_back___rarg(x_24, x_5); +lean_inc(x_7); +x_26 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_26, 0, x_1); +lean_ctor_set(x_26, 1, x_7); +lean_ctor_set(x_26, 2, x_25); +lean_ctor_set(x_3, 0, x_26); +x_27 = lean_box(0); +x_28 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_28, 0, x_5); +lean_ctor_set(x_28, 1, x_6); +lean_ctor_set(x_28, 2, x_7); +lean_ctor_set(x_28, 3, x_3); +lean_ctor_set(x_28, 4, x_27); +return x_28; } } else { +lean_free_object(x_3); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_1); +return x_2; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_3, 1); +lean_inc(x_29); +lean_dec(x_3); +x_30 = lean_array_get_size(x_5); +x_31 = lean_unsigned_to_nat(0u); +x_32 = lean_nat_dec_eq(x_30, x_31); +lean_dec(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + lean_ctor_release(x_2, 3); + lean_ctor_release(x_2, 4); + x_33 = x_2; +} else { + lean_dec_ref(x_2); + x_33 = lean_box(0); +} +x_34 = l_Lean_instInhabitedSyntax; +x_35 = l_Array_back___rarg(x_34, x_5); +lean_inc(x_7); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_1); +lean_ctor_set(x_36, 1, x_7); +lean_ctor_set(x_36, 2, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_29); +x_38 = lean_box(0); +if (lean_is_scalar(x_33)) { + x_39 = lean_alloc_ctor(0, 5, 0); +} else { + x_39 = x_33; +} +lean_ctor_set(x_39, 0, x_5); +lean_ctor_set(x_39, 1, x_6); +lean_ctor_set(x_39, 2, x_7); +lean_ctor_set(x_39, 3, x_37); +lean_ctor_set(x_39, 4, x_38); +return x_39; +} +else +{ +lean_dec(x_29); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_1); return x_2; } } +} else { +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); return x_2; @@ -11151,45 +11413,48 @@ x_7 = lean_string_utf8_at_end(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { -lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_dec(x_1); x_8 = lean_ctor_get(x_3, 3); lean_inc(x_8); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -x_10 = lean_nat_dec_eq(x_9, x_6); -lean_dec(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_dec(x_8); -x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_tokenFnAux(x_2, x_3); -x_12 = l___private_Lean_Parser_Basic_0__Lean_Parser_updateCache(x_6, x_11); -return x_12; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_nat_dec_eq(x_10, x_6); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_9); +x_12 = l___private_Lean_Parser_Basic_0__Lean_Parser_tokenFnAux(x_2, x_3); +x_13 = l___private_Lean_Parser_Basic_0__Lean_Parser_updateTokenCache(x_6, x_12); +return x_13; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_dec(x_6); lean_dec(x_2); -x_13 = lean_ctor_get(x_8, 2); -lean_inc(x_13); -x_14 = l_Lean_Parser_ParserState_pushSyntax(x_3, x_13); -x_15 = lean_ctor_get(x_8, 1); -lean_inc(x_15); -lean_dec(x_8); -x_16 = l_Lean_Parser_ParserState_setPos(x_14, x_15); -return x_16; +x_14 = lean_ctor_get(x_9, 2); +lean_inc(x_14); +x_15 = l_Lean_Parser_ParserState_pushSyntax(x_3, x_14); +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_dec(x_9); +x_17 = l_Lean_Parser_ParserState_setPos(x_15, x_16); +return x_17; } } else { -lean_object* x_17; lean_object* x_18; +lean_object* x_18; lean_object* x_19; lean_dec(x_6); lean_dec(x_2); -x_17 = l_Lean_Parser_ParserState_mkEOIError___closed__1; -x_18 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_17, x_1); -return x_18; +x_18 = l_Lean_Parser_ParserState_mkEOIError___closed__1; +x_19 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_18, x_1); +return x_19; } } } @@ -11240,36 +11505,39 @@ return x_16; LEAN_EXPORT lean_object* l_Lean_Parser_peekToken(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_3 = lean_ctor_get(x_2, 3); lean_inc(x_3); x_4 = lean_ctor_get(x_3, 0); lean_inc(x_4); -x_5 = lean_ctor_get(x_2, 2); +lean_dec(x_3); +x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); -x_6 = lean_nat_dec_eq(x_4, x_5); +x_6 = lean_ctor_get(x_2, 2); +lean_inc(x_6); +x_7 = lean_nat_dec_eq(x_5, x_6); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -if (x_6 == 0) +if (x_7 == 0) { -lean_object* x_7; -lean_dec(x_3); -x_7 = l_Lean_Parser_peekTokenAux(x_1, x_2); -return x_7; +lean_object* x_8; +lean_dec(x_4); +x_8 = l_Lean_Parser_peekTokenAux(x_1, x_2); +return x_8; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_dec(x_1); -x_8 = lean_ctor_get(x_3, 2); -lean_inc(x_8); -lean_dec(x_3); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_10, 1, x_9); -return x_10; +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +lean_dec(x_4); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +return x_11; } } } @@ -11834,8 +12102,8 @@ x_10 = lean_string_utf8_at_end(x_9, x_7); if (x_10 == 0) { uint32_t x_11; uint32_t x_12; uint8_t x_13; -x_11 = lean_string_utf8_get(x_1, x_3); -x_12 = lean_string_utf8_get(x_9, x_7); +x_11 = lean_string_utf8_get_fast(x_1, x_3); +x_12 = lean_string_utf8_get_fast(x_9, x_7); x_13 = lean_uint32_dec_eq(x_11, x_12); if (x_13 == 0) { @@ -11848,9 +12116,9 @@ return x_14; else { lean_object* x_15; lean_object* x_16; -x_15 = lean_string_utf8_next(x_1, x_3); +x_15 = lean_string_utf8_next_fast(x_1, x_3); lean_dec(x_3); -x_16 = l_Lean_Parser_ParserState_next(x_5, x_9, x_7); +x_16 = l_Lean_Parser_ParserState_next_x27(x_5, x_9, x_7, lean_box(0)); lean_dec(x_7); x_3 = x_15; x_5 = x_16; @@ -13372,7 +13640,7 @@ if (x_9 == 0) { lean_object* x_10; uint8_t x_11; x_10 = lean_ctor_get(x_4, 0); -x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(x_3, x_10); +x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(x_3, x_10); if (x_11 == 0) { uint8_t x_12; @@ -13430,7 +13698,7 @@ lean_object* x_23; uint8_t x_24; x_23 = lean_ctor_get(x_4, 0); lean_inc(x_23); lean_dec(x_4); -x_24 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_357_(x_3, x_23); +x_24 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqError____x40_Lean_Parser_Basic___hyg_339_(x_3, x_23); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -24867,7 +25135,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -24879,7 +25147,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; lean_object* x_6; @@ -24887,7 +25155,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586_(x_3, x_4); +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -24896,7 +25164,7 @@ static lean_object* _init_l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10586____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10950____boxed), 2, 0); return x_1; } } @@ -24908,7 +25176,7 @@ x_1 = l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1() { _start: { lean_object* x_1; @@ -24916,17 +25184,17 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.default", 40); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -24935,23 +25203,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -24959,23 +25227,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___elambda__1___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -24983,7 +25251,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8() { _start: { lean_object* x_1; @@ -24991,33 +25259,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.symbol", 39); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -25025,23 +25293,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___elambda__1___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -25049,7 +25317,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14() { _start: { lean_object* x_1; @@ -25057,33 +25325,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.both", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14; x_2 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -25091,23 +25359,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___elambda__1___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15; x_3 = lean_alloc_ctor(3, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18; x_2 = 0; x_3 = lean_alloc_ctor(5, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -25115,7 +25383,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -25127,14 +25395,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5; +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7; +x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -25147,14 +25415,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11; +x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13; +x_13 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -25167,14 +25435,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17; +x_17 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19; +x_19 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -25182,13 +25450,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610_(x_3, x_2); +x_4 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -25197,7 +25465,7 @@ static lean_object* _init_l_Lean_Parser_instReprLeadingIdentBehavior___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____boxed), 2, 0); return x_1; } } @@ -26265,7 +26533,7 @@ lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; @@ -26273,19 +26541,19 @@ x_4 = l_Lean_Parser_whitespace(x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1___boxed), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1; x_3 = lean_st_mk_ref(x_2, x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -26307,16 +26575,16 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____lambda__1(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -26324,11 +26592,11 @@ x_1 = l_Lean_Parser_categoryParserFnRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1; x_3 = lean_st_ref_get(x_2, x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -26350,19 +26618,19 @@ return x_7; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1; x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); return x_3; } @@ -26435,65 +26703,912 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +if (lean_obj_tag(x_2) == 0) { -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_3, 2); -lean_dec(x_6); -lean_ctor_set(x_3, 2, x_2); -x_7 = l_Lean_Parser_categoryParserFn(x_1, x_3, x_4); -return x_7; +uint8_t x_3; +x_3 = 0; +return x_3; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_8 = lean_ctor_get(x_3, 0); -x_9 = lean_ctor_get(x_3, 1); -x_10 = lean_ctor_get(x_3, 3); -x_11 = lean_ctor_get(x_3, 4); -x_12 = lean_ctor_get_uint8(x_3, sizeof(void*)*7); -x_13 = lean_ctor_get(x_3, 5); -x_14 = lean_ctor_get(x_3, 6); +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(x_4, x_1); +if (x_6 == 0) +{ +x_2 = x_5; +goto _start; +} +else +{ +uint8_t x_8; +x_8 = 1; +return x_8; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Parser_categoryParser___elambda__1___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +return x_1; +} +else +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_array_get_size(x_1); +x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_4); +x_8 = (size_t)(x_7) & (lean_unbox(x_6) - 1); +x_9 = lean_array_uget(x_1, x_8); +lean_ctor_set(x_2, 2, x_9); +x_10 = lean_array_uset(x_1, x_8, x_2); +x_1 = x_10; +x_2 = x_5; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint64_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_ctor_get(x_2, 2); lean_inc(x_14); lean_inc(x_13); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_dec(x_3); -x_15 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_15, 0, x_8); -lean_ctor_set(x_15, 1, x_9); -lean_ctor_set(x_15, 2, x_2); -lean_ctor_set(x_15, 3, x_10); -lean_ctor_set(x_15, 4, x_11); -lean_ctor_set(x_15, 5, x_13); -lean_ctor_set(x_15, 6, x_14); -lean_ctor_set_uint8(x_15, sizeof(void*)*7, x_12); -x_16 = l_Lean_Parser_categoryParserFn(x_1, x_15, x_4); -return x_16; +lean_inc(x_12); +lean_dec(x_2); +x_15 = lean_array_get_size(x_1); +x_16 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_12); +x_17 = (size_t)(x_16) & (lean_unbox(x_15) - 1); +x_18 = lean_array_uget(x_1, x_17); +x_19 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 1, x_13); +lean_ctor_set(x_19, 2, x_18); +x_20 = lean_array_uset(x_1, x_17, x_19); +x_1 = x_20; +x_2 = x_14; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Parser_categoryParser___elambda__1___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_3 = lean_alloc_closure((void*)(l_Lean_Parser_categoryParser___elambda__1), 4, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -x_4 = l_Lean_Parser_errorAtSavedPos___closed__2; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); +lean_object* x_4; uint8_t x_5; +x_4 = lean_array_get_size(x_2); +x_5 = lean_nat_dec_lt(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_6 = lean_array_fget(x_2, x_1); +x_7 = lean_box(0); +x_8 = lean_array_fset(x_2, x_1, x_7); +x_9 = l_Lean_AssocList_foldlM___at_Lean_Parser_categoryParser___elambda__1___spec__5(x_3, x_6); +x_10 = lean_unsigned_to_nat(1u); +x_11 = lean_nat_add(x_1, x_10); +lean_dec(x_1); +x_1 = x_11; +x_2 = x_8; +x_3 = x_9; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Parser_categoryParser___elambda__1___spec__3(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_3 = lean_array_get_size(x_2); +x_4 = lean_unsigned_to_nat(2u); +x_5 = lean_nat_mul(x_3, x_4); +lean_dec(x_3); +x_6 = lean_box(0); +x_7 = lean_mk_array(x_5, x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = l_Lean_HashMapImp_moveEntries___at_Lean_Parser_categoryParser___elambda__1___spec__4(x_8, x_2, x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(0); +return x_4; +} +else +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_ctor_get(x_3, 2); +x_9 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(x_6, x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(x_1, x_2, x_8); +lean_ctor_set(x_3, 2, x_10); +return x_3; +} +else +{ +lean_dec(x_7); +lean_dec(x_6); +lean_ctor_set(x_3, 1, x_2); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_3, 0); +x_12 = lean_ctor_get(x_3, 1); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_3); +x_14 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(x_11, x_1); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(x_1, x_2, x_13); +x_16 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_15); +return x_16; +} +else +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_11); +x_17 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_2); +lean_ctor_set(x_17, 2, x_13); +return x_17; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint64_t x_8; size_t x_9; lean_object* x_10; uint8_t x_11; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_array_get_size(x_6); +x_8 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_2); +lean_inc(x_7); +x_9 = (size_t)(x_8) & (lean_unbox(x_7) - 1); +x_10 = lean_array_uget(x_6, x_9); +x_11 = l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2(x_2, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_5, x_12); +lean_dec(x_5); +x_14 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_14, 0, x_2); +lean_ctor_set(x_14, 1, x_3); +lean_ctor_set(x_14, 2, x_10); +x_15 = lean_array_uset(x_6, x_9, x_14); +x_16 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_13); +x_17 = lean_nat_dec_le(x_16, x_7); +lean_dec(x_7); +lean_dec(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_free_object(x_1); +x_18 = l_Lean_HashMapImp_expand___at_Lean_Parser_categoryParser___elambda__1___spec__3(x_13, x_15); +return x_18; +} +else +{ +lean_ctor_set(x_1, 1, x_15); +lean_ctor_set(x_1, 0, x_13); +return x_1; +} +} +else +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_7); +x_19 = l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(x_2, x_3, x_10); +x_20 = lean_array_uset(x_6, x_9, x_19); +lean_ctor_set(x_1, 1, x_20); +return x_1; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint64_t x_24; size_t x_25; lean_object* x_26; uint8_t x_27; +x_21 = lean_ctor_get(x_1, 0); +x_22 = lean_ctor_get(x_1, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_1); +x_23 = lean_array_get_size(x_22); +x_24 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_2); +lean_inc(x_23); +x_25 = (size_t)(x_24) & (lean_unbox(x_23) - 1); +x_26 = lean_array_uget(x_22, x_25); +x_27 = l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2(x_2, x_26); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_28 = lean_unsigned_to_nat(1u); +x_29 = lean_nat_add(x_21, x_28); +lean_dec(x_21); +x_30 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_30, 0, x_2); +lean_ctor_set(x_30, 1, x_3); +lean_ctor_set(x_30, 2, x_26); +x_31 = lean_array_uset(x_22, x_25, x_30); +x_32 = l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(x_29); +x_33 = lean_nat_dec_le(x_32, x_23); +lean_dec(x_23); +lean_dec(x_32); +if (x_33 == 0) +{ +lean_object* x_34; +x_34 = l_Lean_HashMapImp_expand___at_Lean_Parser_categoryParser___elambda__1___spec__3(x_29, x_31); +return x_34; +} +else +{ +lean_object* x_35; +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_29); +lean_ctor_set(x_35, 1, x_31); +return x_35; +} +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_23); +x_36 = l_Lean_AssocList_replace___at_Lean_Parser_categoryParser___elambda__1___spec__6(x_2, x_3, x_26); +x_37 = lean_array_uset(x_22, x_25, x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_21); +lean_ctor_set(x_38, 1, x_37); +return x_38; +} +} +} +} +static lean_object* _init_l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Id_instMonadId; +x_2 = l_instInhabitedPUnit; +x_3 = l_instInhabited___rarg(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_3; +x_3 = lean_box(0); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = lean_ctor_get(x_2, 2); +x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqCategoryCacheKey____x40_Lean_Parser_Basic___hyg_135_(x_4, x_1); +if (x_7 == 0) +{ +x_2 = x_6; +goto _start; +} +else +{ +lean_object* x_9; +lean_inc(x_5); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_5); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 1); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_array_get_size(x_3); +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_hashCategoryCacheKey____x40_Lean_Parser_Basic___hyg_246_(x_2); +x_6 = (size_t)(x_5) & (lean_unbox(x_4) - 1); +x_7 = lean_array_uget(x_3, x_6); +lean_dec(x_3); +x_8 = l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9(x_2, x_7); +lean_dec(x_7); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; uint8_t x_6; +x_5 = l_Lean_Parser_ParserState_stackSize(x_1); +x_6 = lean_nat_dec_eq(x_5, x_2); +lean_dec(x_5); +if (x_6 == 0) +{ +uint8_t x_7; +x_7 = !lean_is_exclusive(x_1); +if (x_7 == 0) +{ +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_1, 3); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 2); +x_13 = lean_ctor_get(x_1, 4); +x_14 = lean_ctor_get(x_8, 1); +x_15 = l_Lean_instInhabitedSyntax; +x_16 = l_Array_back___rarg(x_15, x_10); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +x_17 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_11); +lean_ctor_set(x_17, 2, x_12); +lean_ctor_set(x_17, 3, x_13); +x_18 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_14, x_3, x_17); +lean_ctor_set(x_8, 1, x_18); +return x_1; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_19 = lean_ctor_get(x_1, 0); +x_20 = lean_ctor_get(x_1, 1); +x_21 = lean_ctor_get(x_1, 2); +x_22 = lean_ctor_get(x_1, 4); +x_23 = lean_ctor_get(x_8, 0); +x_24 = lean_ctor_get(x_8, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_8); +x_25 = l_Lean_instInhabitedSyntax; +x_26 = l_Array_back___rarg(x_25, x_19); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_20); +lean_ctor_set(x_27, 2, x_21); +lean_ctor_set(x_27, 3, x_22); +x_28 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_24, x_3, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_23); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_1, 3, x_29); +return x_1; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_30 = lean_ctor_get(x_1, 3); +x_31 = lean_ctor_get(x_1, 0); +x_32 = lean_ctor_get(x_1, 1); +x_33 = lean_ctor_get(x_1, 2); +x_34 = lean_ctor_get(x_1, 4); +lean_inc(x_34); +lean_inc(x_30); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_1); +x_35 = lean_ctor_get(x_30, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_30, 1); +lean_inc(x_36); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_37 = x_30; +} else { + lean_dec_ref(x_30); + x_37 = lean_box(0); +} +x_38 = l_Lean_instInhabitedSyntax; +x_39 = l_Array_back___rarg(x_38, x_31); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +x_40 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_32); +lean_ctor_set(x_40, 2, x_33); +lean_ctor_set(x_40, 3, x_34); +x_41 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_36, x_3, x_40); +if (lean_is_scalar(x_37)) { + x_42 = lean_alloc_ctor(0, 2, 0); +} else { + x_42 = x_37; +} +lean_ctor_set(x_42, 0, x_35); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_43, 0, x_31); +lean_ctor_set(x_43, 1, x_32); +lean_ctor_set(x_43, 2, x_33); +lean_ctor_set(x_43, 3, x_42); +lean_ctor_set(x_43, 4, x_34); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_44 = lean_box(0); +lean_inc(x_1); +x_45 = l_Lean_Parser_ParserState_pushSyntax(x_1, x_44); +x_46 = lean_ctor_get(x_1, 3); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = !lean_is_exclusive(x_1); +if (x_48 == 0) +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_49 = lean_ctor_get(x_1, 1); +x_50 = lean_ctor_get(x_1, 2); +x_51 = lean_ctor_get(x_1, 4); +x_52 = lean_ctor_get(x_1, 3); +lean_dec(x_52); +x_53 = lean_ctor_get(x_1, 0); +lean_dec(x_53); +x_54 = !lean_is_exclusive(x_46); +if (x_54 == 0) +{ +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_55 = lean_ctor_get(x_46, 1); +x_56 = l_Lean_instInhabitedSyntax; +x_57 = l_Array_back___rarg(x_56, x_47); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +x_58 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_49); +lean_ctor_set(x_58, 2, x_50); +lean_ctor_set(x_58, 3, x_51); +x_59 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_55, x_3, x_58); +lean_ctor_set(x_46, 1, x_59); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_60 = lean_ctor_get(x_46, 0); +x_61 = lean_ctor_get(x_46, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_46); +x_62 = l_Lean_instInhabitedSyntax; +x_63 = l_Array_back___rarg(x_62, x_47); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +x_64 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_49); +lean_ctor_set(x_64, 2, x_50); +lean_ctor_set(x_64, 3, x_51); +x_65 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_61, x_3, x_64); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_60); +lean_ctor_set(x_66, 1, x_65); +lean_ctor_set(x_1, 3, x_66); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} +} +else +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_67 = lean_ctor_get(x_1, 1); +x_68 = lean_ctor_get(x_1, 2); +x_69 = lean_ctor_get(x_1, 4); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_1); +x_70 = lean_ctor_get(x_46, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_46, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_72 = x_46; +} else { + lean_dec_ref(x_46); + x_72 = lean_box(0); +} +x_73 = l_Lean_instInhabitedSyntax; +x_74 = l_Array_back___rarg(x_73, x_47); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +x_75 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_67); +lean_ctor_set(x_75, 2, x_68); +lean_ctor_set(x_75, 3, x_69); +x_76 = l_Lean_HashMap_insert___at_Lean_Parser_categoryParser___elambda__1___spec__1(x_71, x_3, x_75); +if (lean_is_scalar(x_72)) { + x_77 = lean_alloc_ctor(0, 2, 0); +} else { + x_77 = x_72; +} +lean_ctor_set(x_77, 0, x_70); +lean_ctor_set(x_77, 1, x_76); +x_78 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_78, 0, x_47); +lean_ctor_set(x_78, 1, x_67); +lean_ctor_set(x_78, 2, x_68); +lean_ctor_set(x_78, 3, x_77); +lean_ctor_set(x_78, 4, x_69); +return x_78; +} +} +} +} +static lean_object* _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Parser.Basic", 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Parser.categoryParser", 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("categoryParser: unexpected stack growth", 39); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1; +x_2 = l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2; +x_3 = lean_unsigned_to_nat(1711u); +x_4 = lean_unsigned_to_nat(6u); +x_5 = l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = l_Lean_Parser_ParserState_stackSize(x_1); +x_8 = !lean_is_exclusive(x_2); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_9 = lean_ctor_get(x_2, 2); +lean_dec(x_9); +lean_ctor_set(x_2, 2, x_3); +x_10 = l_Lean_Parser_categoryParserFn(x_4, x_2, x_1); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_7, x_11); +x_13 = l_Lean_Parser_ParserState_stackSize(x_10); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_box(0); +x_16 = l_Lean_Parser_categoryParser___elambda__1___lambda__1(x_10, x_7, x_5, x_15); +lean_dec(x_7); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4; +x_18 = l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7(x_17); +x_19 = l_Lean_Parser_categoryParser___elambda__1___lambda__1(x_10, x_7, x_5, x_18); +lean_dec(x_18); +lean_dec(x_7); +return x_19; +} +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_20 = lean_ctor_get(x_2, 0); +x_21 = lean_ctor_get(x_2, 1); +x_22 = lean_ctor_get(x_2, 3); +x_23 = lean_ctor_get(x_2, 4); +x_24 = lean_ctor_get_uint8(x_2, sizeof(void*)*7); +x_25 = lean_ctor_get(x_2, 5); +x_26 = lean_ctor_get(x_2, 6); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_2); +x_27 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_21); +lean_ctor_set(x_27, 2, x_3); +lean_ctor_set(x_27, 3, x_22); +lean_ctor_set(x_27, 4, x_23); +lean_ctor_set(x_27, 5, x_25); +lean_ctor_set(x_27, 6, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*7, x_24); +x_28 = l_Lean_Parser_categoryParserFn(x_4, x_27, x_1); +x_29 = lean_unsigned_to_nat(1u); +x_30 = lean_nat_add(x_7, x_29); +x_31 = l_Lean_Parser_ParserState_stackSize(x_28); +x_32 = lean_nat_dec_lt(x_30, x_31); +lean_dec(x_31); +lean_dec(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_box(0); +x_34 = l_Lean_Parser_categoryParser___elambda__1___lambda__1(x_28, x_7, x_5, x_33); +lean_dec(x_7); +return x_34; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4; +x_36 = l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7(x_35); +x_37 = l_Lean_Parser_categoryParser___elambda__1___lambda__1(x_28, x_7, x_5, x_36); +lean_dec(x_36); +lean_dec(x_7); +return x_37; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 2); +lean_inc(x_6); +x_7 = lean_ctor_get(x_4, 3); +lean_inc(x_7); +lean_inc(x_2); +lean_inc(x_1); +x_8 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_2); +lean_ctor_set(x_8, 2, x_6); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_inc(x_8); +x_10 = l_Lean_HashMapImp_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__8(x_9, x_8); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_7); +lean_dec(x_5); +x_11 = lean_box(0); +x_12 = l_Lean_Parser_categoryParser___elambda__1___lambda__2(x_4, x_3, x_2, x_1, x_8, x_11); +return x_12; +} +else +{ +uint8_t x_13; +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_4); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_14 = lean_ctor_get(x_4, 4); +lean_dec(x_14); +x_15 = lean_ctor_get(x_4, 3); +lean_dec(x_15); +x_16 = lean_ctor_get(x_4, 2); +lean_dec(x_16); +x_17 = lean_ctor_get(x_4, 1); +lean_dec(x_17); +x_18 = lean_ctor_get(x_4, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_10, 0); +lean_inc(x_19); +lean_dec(x_10); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 2); +lean_inc(x_22); +x_23 = lean_ctor_get(x_19, 3); +lean_inc(x_23); +lean_dec(x_19); +x_24 = lean_array_push(x_5, x_20); +lean_ctor_set(x_4, 4, x_23); +lean_ctor_set(x_4, 2, x_22); +lean_ctor_set(x_4, 1, x_21); +lean_ctor_set(x_4, 0, x_24); +return x_4; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_4); +x_25 = lean_ctor_get(x_10, 0); +lean_inc(x_25); +lean_dec(x_10); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_25, 2); +lean_inc(x_28); +x_29 = lean_ctor_get(x_25, 3); +lean_inc(x_29); +lean_dec(x_25); +x_30 = lean_array_push(x_5, x_26); +x_31 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +lean_ctor_set(x_31, 2, x_28); +lean_ctor_set(x_31, 3, x_7); +lean_ctor_set(x_31, 4, x_29); +return x_31; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_3 = lean_alloc_closure((void*)(l_Lean_Parser_categoryParser___elambda__1), 4, 2); +lean_closure_set(x_3, 0, x_1); +lean_closure_set(x_3, 1, x_2); +x_4 = l_Lean_Parser_errorAtSavedPos___closed__2; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_AssocList_contains___at_Lean_Parser_categoryParser___elambda__1___spec__2(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_AssocList_find_x3f___at_Lean_Parser_categoryParser___elambda__1___spec__9(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Parser_categoryParser___elambda__1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_4); +lean_dec(x_2); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Parser_categoryParser___elambda__1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_6); +return x_7; +} +} static lean_object* _init_l_Lean_Parser_termParser___closed__1() { _start: { @@ -26555,7 +27670,7 @@ x_10 = lean_string_utf8_at_end(x_8, x_9); if (x_10 == 0) { uint32_t x_11; uint32_t x_12; uint8_t x_13; -x_11 = lean_string_utf8_get(x_8, x_9); +x_11 = lean_string_utf8_get_fast(x_8, x_9); lean_dec(x_9); x_12 = 58; x_13 = lean_uint32_dec_eq(x_11, x_12); @@ -27281,21 +28396,20 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser(lean_object* x_1) { +static lean_object* _init_l_Lean_Parser_instCoeStringParser___closed__1() { _start: { -lean_object* x_2; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_symbol___boxed), 1, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Parser_instCoeStringParser() { _start: { -lean_object* x_2; -x_2 = l_Lean_Parser_instCoeStringParser(x_1); -lean_dec(x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_Parser_instCoeStringParser___closed__1; +return x_1; } } LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol(lean_object* x_1, uint8_t x_2) { @@ -27421,7 +28535,7 @@ static lean_object* _init_l_Lean_Parser_mkAntiquot___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_instInhabitedInputContext___closed__2; +x_1 = l_Lean_Parser_Error_unexpected___default___closed__1; x_2 = l_Lean_Parser_checkNoWsBefore(x_1); return x_2; } @@ -28089,48 +29203,46 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_unsigned_to_nat(1u); x_7 = lean_nat_add(x_1, x_6); x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); if (x_8 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_nat_sub(x_5, x_1); +lean_object* x_9; lean_object* x_10; +x_9 = lean_nat_sub(x_5, x_7); +lean_dec(x_7); lean_dec(x_5); -x_10 = lean_nat_sub(x_9, x_6); +x_10 = lean_array_fget(x_4, x_9); lean_dec(x_9); -x_11 = l_Lean_instInhabitedSyntax; -x_12 = lean_array_get(x_11, x_4, x_10); -lean_dec(x_10); lean_dec(x_4); -if (lean_obj_tag(x_12) == 3) +if (lean_obj_tag(x_10) == 3) { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_12, 2); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_Parser_categoryParserFn(x_13, x_2, x_3); -return x_14; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_10, 2); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Parser_categoryParserFn(x_11, x_2, x_3); +return x_12; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_12); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_10); lean_dec(x_2); -x_15 = lean_box(0); -x_16 = l_Lean_Parser_categoryParserOfStackFn___closed__1; -x_17 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_16, x_15); -return x_17; +x_13 = lean_box(0); +x_14 = l_Lean_Parser_categoryParserOfStackFn___closed__1; +x_15 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_14, x_13); +return x_15; } } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_18 = lean_box(0); -x_19 = l_Lean_Parser_categoryParserOfStackFn___closed__2; -x_20 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_19, x_18); -return x_20; +x_16 = lean_box(0); +x_17 = l_Lean_Parser_categoryParserOfStackFn___closed__2; +x_18 = l_Lean_Parser_ParserState_mkUnexpectedError(x_3, x_17, x_16); +return x_18; } } } @@ -28776,119 +29888,119 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___ size_t x_7; size_t x_8; lean_object* x_9; x_7 = 1; x_8 = lean_usize_add(x_1, x_7); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_2, x_3, x_4, x_8, x_5, x_6); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_2, lean_box(0), x_3, x_4, x_8, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_5, x_6); +if (x_8 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -x_9 = lean_array_uget(x_3, x_4); -lean_inc(x_2); -x_10 = lean_apply_2(x_2, x_9, x_6); -x_11 = lean_box_usize(x_4); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +x_10 = lean_array_uget(x_4, x_5); +lean_inc(x_3); +x_11 = lean_apply_2(x_3, x_10, x_7); x_12 = lean_box_usize(x_5); -x_13 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___lambda__1___boxed), 6, 5); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_1); -lean_closure_set(x_13, 2, x_2); -lean_closure_set(x_13, 3, x_3); -lean_closure_set(x_13, 4, x_12); -x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_10, x_13); -return x_14; +x_13 = lean_box_usize(x_6); +x_14 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_1); +lean_closure_set(x_14, 2, x_3); +lean_closure_set(x_14, 3, x_4); +lean_closure_set(x_14, 4, x_13); +x_15 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_11, x_14); +return x_15; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_15 = lean_ctor_get(x_1, 0); -lean_inc(x_15); -lean_dec(x_1); -x_16 = lean_ctor_get(x_15, 1); +x_16 = lean_ctor_get(x_1, 0); lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_apply_2(x_16, lean_box(0), x_6); -return x_17; +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_apply_2(x_17, lean_box(0), x_7); +return x_18; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed), 6, 0); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed), 7, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = l_Lean_Syntax_getArgs(x_2); -lean_dec(x_2); -x_6 = lean_array_get_size(x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_lt(x_7, x_6); -if (x_8 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = l_Lean_Syntax_getArgs(x_3); +lean_dec(x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +if (x_9 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); +lean_dec(x_4); +x_10 = lean_ctor_get(x_1, 0); lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_4); -return x_11; +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_5); +return x_12; } else { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_6, x_6); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_nat_dec_le(x_7, x_7); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_13, 1); +lean_dec(x_4); +x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_apply_2(x_14, lean_box(0), x_4); -return x_15; +lean_dec(x_1); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_apply_2(x_15, lean_box(0), x_5); +return x_16; } else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = 0; -x_17 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_1, x_3, x_5, x_16, x_17, x_4); -return x_18; +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = 0; +x_18 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_19 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_1, lean_box(0), x_4, x_6, x_17, x_18, x_5); +return x_19; } } } } -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Syntax_foldArgsM___rarg), 4, 0); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_foldArgsM___rarg), 5, 0); +return x_2; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { @@ -28903,16 +30015,16 @@ x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___la return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); +size_t x_8; size_t x_9; lean_object* x_10; x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_1, x_2, x_3, x_7, x_8, x_6); -return x_9; +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +return x_10; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgs___spec__2___rarg(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { @@ -29030,47 +30142,47 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___s size_t x_7; size_t x_8; lean_object* x_9; x_7 = 1; x_8 = lean_usize_add(x_1, x_7); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_2, x_3, x_4, x_8, x_5, x_6); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_2, lean_box(0), x_3, x_4, x_8, x_5, x_6); return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_eq(x_4, x_5); -if (x_7 == 0) +uint8_t x_8; +x_8 = lean_usize_dec_eq(x_5, x_6); +if (x_8 == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_8 = lean_ctor_get(x_1, 1); -lean_inc(x_8); -x_9 = lean_array_uget(x_3, x_4); -lean_inc(x_2); -x_10 = lean_apply_2(x_2, x_9, x_6); -x_11 = lean_box_usize(x_4); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +x_10 = lean_array_uget(x_4, x_5); +lean_inc(x_3); +x_11 = lean_apply_2(x_3, x_10, x_7); x_12 = lean_box_usize(x_5); -x_13 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1___boxed), 6, 5); -lean_closure_set(x_13, 0, x_11); -lean_closure_set(x_13, 1, x_1); -lean_closure_set(x_13, 2, x_2); -lean_closure_set(x_13, 3, x_3); -lean_closure_set(x_13, 4, x_12); -x_14 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_10, x_13); -return x_14; +x_13 = lean_box_usize(x_6); +x_14 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1___boxed), 6, 5); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_1); +lean_closure_set(x_14, 2, x_3); +lean_closure_set(x_14, 3, x_4); +lean_closure_set(x_14, 4, x_13); +x_15 = lean_apply_4(x_9, lean_box(0), lean_box(0), x_11, x_14); +return x_15; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_15 = lean_ctor_get(x_1, 0); -lean_inc(x_15); -lean_dec(x_1); -x_16 = lean_ctor_get(x_15, 1); +x_16 = lean_ctor_get(x_1, 0); lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_apply_2(x_16, lean_box(0), x_6); -return x_17; +lean_dec(x_1); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_apply_2(x_17, lean_box(0), x_7); +return x_18; } } } @@ -29078,61 +30190,61 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___s _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed), 6, 0); +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed), 7, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = l_Lean_Syntax_getArgs(x_2); -lean_dec(x_2); -x_6 = lean_array_get_size(x_5); -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_lt(x_7, x_6); -if (x_8 == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_6 = l_Lean_Syntax_getArgs(x_3); +lean_dec(x_3); +x_7 = lean_array_get_size(x_6); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +if (x_9 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -lean_dec(x_1); -x_10 = lean_ctor_get(x_9, 1); +lean_dec(x_4); +x_10 = lean_ctor_get(x_1, 0); lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_apply_2(x_10, lean_box(0), x_4); -return x_11; +lean_dec(x_1); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_apply_2(x_11, lean_box(0), x_5); +return x_12; } else { -uint8_t x_12; -x_12 = lean_nat_dec_le(x_6, x_6); -if (x_12 == 0) +uint8_t x_13; +x_13 = lean_nat_dec_le(x_7, x_7); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -x_14 = lean_ctor_get(x_13, 1); +lean_dec(x_4); +x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_apply_2(x_14, lean_box(0), x_4); -return x_15; +lean_dec(x_1); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_apply_2(x_15, lean_box(0), x_5); +return x_16; } else { -size_t x_16; size_t x_17; lean_object* x_18; -x_16 = 0; -x_17 = lean_usize_of_nat(x_6); -lean_dec(x_6); -x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_1, x_3, x_5, x_16, x_17, x_4); -return x_18; +size_t x_17; size_t x_18; lean_object* x_19; +x_17 = 0; +x_18 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_19 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_1, lean_box(0), x_4, x_6, x_17, x_18, x_5); +return x_19; } } } @@ -29141,7 +30253,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spe _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg), 5, 0); return x_2; } } @@ -29160,7 +30272,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_alloc_closure((void*)(l_Lean_Syntax_forArgsM___rarg___lambda__1___boxed), 3, 1); lean_closure_set(x_4, 0, x_3); x_5 = lean_box(0); -x_6 = l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(x_1, x_2, x_4, x_5); +x_6 = l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(x_1, lean_box(0), x_2, x_4, x_5); return x_6; } } @@ -29184,16 +30296,16 @@ x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lam return x_9; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -size_t x_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); +size_t x_8; size_t x_9; lean_object* x_10; x_8 = lean_unbox_usize(x_5); lean_dec(x_5); -x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_1, x_2, x_3, x_7, x_8, x_6); -return x_9; +x_9 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -29240,6 +30352,28 @@ l_Lean_Parser_TokenCacheEntry_stopPos___default = _init_l_Lean_Parser_TokenCache lean_mark_persistent(l_Lean_Parser_TokenCacheEntry_stopPos___default); l_Lean_Parser_TokenCacheEntry_token___default = _init_l_Lean_Parser_TokenCacheEntry_token___default(); lean_mark_persistent(l_Lean_Parser_TokenCacheEntry_token___default); +l_Lean_Parser_instBEqCategoryCacheKey___closed__1 = _init_l_Lean_Parser_instBEqCategoryCacheKey___closed__1(); +lean_mark_persistent(l_Lean_Parser_instBEqCategoryCacheKey___closed__1); +l_Lean_Parser_instBEqCategoryCacheKey = _init_l_Lean_Parser_instBEqCategoryCacheKey(); +lean_mark_persistent(l_Lean_Parser_instBEqCategoryCacheKey); +l_Lean_Parser_instHashableCategoryCacheKey___closed__1 = _init_l_Lean_Parser_instHashableCategoryCacheKey___closed__1(); +lean_mark_persistent(l_Lean_Parser_instHashableCategoryCacheKey___closed__1); +l_Lean_Parser_instHashableCategoryCacheKey = _init_l_Lean_Parser_instHashableCategoryCacheKey(); +lean_mark_persistent(l_Lean_Parser_instHashableCategoryCacheKey); +l_Lean_Parser_Error_unexpected___default___closed__1 = _init_l_Lean_Parser_Error_unexpected___default___closed__1(); +lean_mark_persistent(l_Lean_Parser_Error_unexpected___default___closed__1); +l_Lean_Parser_Error_unexpected___default = _init_l_Lean_Parser_Error_unexpected___default(); +lean_mark_persistent(l_Lean_Parser_Error_unexpected___default); +l_Lean_Parser_Error_expected___default = _init_l_Lean_Parser_Error_expected___default(); +lean_mark_persistent(l_Lean_Parser_Error_expected___default); +l_Lean_Parser_instInhabitedError___closed__1 = _init_l_Lean_Parser_instInhabitedError___closed__1(); +lean_mark_persistent(l_Lean_Parser_instInhabitedError___closed__1); +l_Lean_Parser_instInhabitedError = _init_l_Lean_Parser_instInhabitedError(); +lean_mark_persistent(l_Lean_Parser_instInhabitedError); +l_Lean_Parser_instBEqError___closed__1 = _init_l_Lean_Parser_instBEqError___closed__1(); +lean_mark_persistent(l_Lean_Parser_instBEqError___closed__1); +l_Lean_Parser_instBEqError = _init_l_Lean_Parser_instBEqError(); +lean_mark_persistent(l_Lean_Parser_instBEqError); l_Lean_Parser_initCacheForInput___closed__1 = _init_l_Lean_Parser_initCacheForInput___closed__1(); lean_mark_persistent(l_Lean_Parser_initCacheForInput___closed__1); l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2___closed__1(); @@ -29252,8 +30386,6 @@ l_Lean_Parser_instInhabitedInputContext___closed__2 = _init_l_Lean_Parser_instIn lean_mark_persistent(l_Lean_Parser_instInhabitedInputContext___closed__2); l_Lean_Parser_instInhabitedInputContext___closed__3 = _init_l_Lean_Parser_instInhabitedInputContext___closed__3(); lean_mark_persistent(l_Lean_Parser_instInhabitedInputContext___closed__3); -l_Lean_Parser_instInhabitedInputContext___closed__4 = _init_l_Lean_Parser_instInhabitedInputContext___closed__4(); -lean_mark_persistent(l_Lean_Parser_instInhabitedInputContext___closed__4); l_Lean_Parser_instInhabitedInputContext = _init_l_Lean_Parser_instInhabitedInputContext(); lean_mark_persistent(l_Lean_Parser_instInhabitedInputContext); l_Lean_Parser_ParserModuleContext_currNamespace___default = _init_l_Lean_Parser_ParserModuleContext_currNamespace___default(); @@ -29267,18 +30399,6 @@ l_Lean_Parser_ParserContext_savedPos_x3f___default = _init_l_Lean_Parser_ParserC lean_mark_persistent(l_Lean_Parser_ParserContext_savedPos_x3f___default); l_Lean_Parser_ParserContext_forbiddenTk_x3f___default = _init_l_Lean_Parser_ParserContext_forbiddenTk_x3f___default(); lean_mark_persistent(l_Lean_Parser_ParserContext_forbiddenTk_x3f___default); -l_Lean_Parser_Error_unexpected___default = _init_l_Lean_Parser_Error_unexpected___default(); -lean_mark_persistent(l_Lean_Parser_Error_unexpected___default); -l_Lean_Parser_Error_expected___default = _init_l_Lean_Parser_Error_expected___default(); -lean_mark_persistent(l_Lean_Parser_Error_expected___default); -l_Lean_Parser_instInhabitedError___closed__1 = _init_l_Lean_Parser_instInhabitedError___closed__1(); -lean_mark_persistent(l_Lean_Parser_instInhabitedError___closed__1); -l_Lean_Parser_instInhabitedError = _init_l_Lean_Parser_instInhabitedError(); -lean_mark_persistent(l_Lean_Parser_instInhabitedError); -l_Lean_Parser_instBEqError___closed__1 = _init_l_Lean_Parser_instBEqError___closed__1(); -lean_mark_persistent(l_Lean_Parser_instBEqError___closed__1); -l_Lean_Parser_instBEqError = _init_l_Lean_Parser_instBEqError(); -lean_mark_persistent(l_Lean_Parser_instBEqError); l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__1(); lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__1); l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__2 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_Error_expectedToString___closed__2(); @@ -29479,12 +30599,12 @@ l_Lean_Parser_mkTokenAndFixPos___closed__1 = _init_l_Lean_Parser_mkTokenAndFixPo lean_mark_persistent(l_Lean_Parser_mkTokenAndFixPos___closed__1); l_Lean_Parser_mkTokenAndFixPos___closed__2 = _init_l_Lean_Parser_mkTokenAndFixPos___closed__2(); lean_mark_persistent(l_Lean_Parser_mkTokenAndFixPos___closed__2); -l_Lean_Parser_identFnAux_parse___lambda__2___closed__1 = _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Parser_identFnAux_parse___lambda__2___closed__1); -l_Lean_Parser_identFnAux_parse___lambda__2___closed__2 = _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Parser_identFnAux_parse___lambda__2___closed__2); -l_Lean_Parser_identFnAux_parse___lambda__2___closed__3 = _init_l_Lean_Parser_identFnAux_parse___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Parser_identFnAux_parse___lambda__2___closed__3); +l_Lean_Parser_identFnAux_parse___closed__1 = _init_l_Lean_Parser_identFnAux_parse___closed__1(); +lean_mark_persistent(l_Lean_Parser_identFnAux_parse___closed__1); +l_Lean_Parser_identFnAux_parse___closed__2 = _init_l_Lean_Parser_identFnAux_parse___closed__2(); +lean_mark_persistent(l_Lean_Parser_identFnAux_parse___closed__2); +l_Lean_Parser_identFnAux_parse___closed__3 = _init_l_Lean_Parser_identFnAux_parse___closed__3(); +lean_mark_persistent(l_Lean_Parser_identFnAux_parse___closed__3); l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux___closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux___closed__1(); lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_nameLitAux___closed__1); l_Lean_Parser_symbolInfo___closed__1 = _init_l_Lean_Parser_symbolInfo___closed__1(); @@ -29628,44 +30748,44 @@ l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1 = _init_l_Lean_Parser_inst lean_mark_persistent(l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1); l_Lean_Parser_instBEqLeadingIdentBehavior = _init_l_Lean_Parser_instBEqLeadingIdentBehavior(); lean_mark_persistent(l_Lean_Parser_instBEqLeadingIdentBehavior); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__1); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__2); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__3); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__4); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__5); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__6); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__7); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__8); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__9); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__10); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__11); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__12); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__13); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__14); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__15); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__16); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__17); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__18); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10610____closed__19); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__1); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__2); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__3); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__4); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__5); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__6); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__7); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__8); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__9); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__10); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__11); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__12); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__13); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__14); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__15); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__16); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__17); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__18); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_10974____closed__19); l_Lean_Parser_instReprLeadingIdentBehavior___closed__1 = _init_l_Lean_Parser_instReprLeadingIdentBehavior___closed__1(); lean_mark_persistent(l_Lean_Parser_instReprLeadingIdentBehavior___closed__1); l_Lean_Parser_instReprLeadingIdentBehavior = _init_l_Lean_Parser_instReprLeadingIdentBehavior(); @@ -29680,18 +30800,18 @@ l_Lean_Parser_instInhabitedParserCategory = _init_l_Lean_Parser_instInhabitedPar lean_mark_persistent(l_Lean_Parser_instInhabitedParserCategory); l_Lean_Parser_indexed___rarg___closed__1 = _init_l_Lean_Parser_indexed___rarg___closed__1(); lean_mark_persistent(l_Lean_Parser_indexed___rarg___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141____closed__1); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11141_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492____closed__1); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11492_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_categoryParserFnRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_categoryParserFnRef); lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____lambda__1___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179____closed__1); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11179_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____lambda__1___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530____closed__1); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_11530_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_categoryParserFnExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_categoryParserFnExtension); @@ -29704,6 +30824,16 @@ l_Lean_Parser_categoryParserFn___closed__2 = _init_l_Lean_Parser_categoryParserF lean_mark_persistent(l_Lean_Parser_categoryParserFn___closed__2); l_Lean_Parser_categoryParserFn___closed__3 = _init_l_Lean_Parser_categoryParserFn___closed__3(); lean_mark_persistent(l_Lean_Parser_categoryParserFn___closed__3); +l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1 = _init_l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Parser_categoryParser___elambda__1___spec__7___closed__1); +l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1 = _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__1); +l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2 = _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__2); +l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3 = _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__3); +l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4 = _init_l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Parser_categoryParser___elambda__1___lambda__2___closed__4); l_Lean_Parser_termParser___closed__1 = _init_l_Lean_Parser_termParser___closed__1(); lean_mark_persistent(l_Lean_Parser_termParser___closed__1); l_Lean_Parser_termParser___closed__2 = _init_l_Lean_Parser_termParser___closed__2(); @@ -29780,6 +30910,10 @@ l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__9 = _init_l_Lean_Parser_toke lean_mark_persistent(l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__9); l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__10 = _init_l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__10(); lean_mark_persistent(l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__10); +l_Lean_Parser_instCoeStringParser___closed__1 = _init_l_Lean_Parser_instCoeStringParser___closed__1(); +lean_mark_persistent(l_Lean_Parser_instCoeStringParser___closed__1); +l_Lean_Parser_instCoeStringParser = _init_l_Lean_Parser_instCoeStringParser(); +lean_mark_persistent(l_Lean_Parser_instCoeStringParser); l_Lean_Parser_mkAntiquot___closed__1 = _init_l_Lean_Parser_mkAntiquot___closed__1(); lean_mark_persistent(l_Lean_Parser_mkAntiquot___closed__1); l_Lean_Parser_mkAntiquot___closed__2 = _init_l_Lean_Parser_mkAntiquot___closed__2(); diff --git a/tests/lean/1021.lean.expected.out b/tests/lean/1021.lean.expected.out index 12fdf9d22a7f..9983172c9430 100644 --- a/tests/lean/1021.lean.expected.out +++ b/tests/lean/1021.lean.expected.out @@ -1,8 +1,8 @@ -some { range := { pos := { line := 128, column := 42 }, +some { range := { pos := { line := 133, column := 42 }, charUtf16 := 42, - endPos := { line := 134, column := 31 }, + endPos := { line := 139, column := 31 }, endCharUtf16 := 31 }, - selectionRange := { pos := { line := 128, column := 46 }, + selectionRange := { pos := { line := 133, column := 46 }, charUtf16 := 46, - endPos := { line := 128, column := 58 }, + endPos := { line := 133, column := 58 }, endCharUtf16 := 58 } } diff --git a/tests/lean/1760.lean b/tests/lean/1760.lean new file mode 100644 index 000000000000..806a155dccab --- /dev/null +++ b/tests/lean/1760.lean @@ -0,0 +1,3 @@ +-- should run in linear time +#check [[[[[[[[[[[[[[[[[[[[[[[[ +#check (((((((((((((((((((((((()))))))))))))))))))))))) diff --git a/tests/lean/1760.lean.expected.out b/tests/lean/1760.lean.expected.out new file mode 100644 index 000000000000..0472ebebd019 --- /dev/null +++ b/tests/lean/1760.lean.expected.out @@ -0,0 +1,2 @@ +1760.lean:3:0: error: expected ':', ']' or term +() : Unit diff --git a/tests/lean/noTabs.lean.expected.out b/tests/lean/noTabs.lean.expected.out index 6037902e5c4d..9e301c50810f 100644 --- a/tests/lean/noTabs.lean.expected.out +++ b/tests/lean/noTabs.lean.expected.out @@ -1 +1,3 @@ noTabs.lean:3:0: error: tabs are not allowed; please configure your editor to expand them +let a := 1; +sorryAx ?m true : ?m